author | haftmann |
Sat, 19 Jul 2025 18:41:55 +0200 | |
changeset 82886 | 8d1e295aab70 |
parent 82383 | 7ed32d7f8317 |
permissions | -rw-r--r-- |
76183 | 1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_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 |
78788 | 4 |
Author: Martin Desharnais, MPI-INF Saarbruecken |
28592 | 5 |
|
36376 | 6 |
Setup for supported ATPs. |
28592 | 7 |
*) |
8 |
||
72400 | 9 |
signature SLEDGEHAMMER_ATP_SYSTEMS = |
28592 | 10 |
sig |
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 |
|
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
76301
diff
changeset
|
15 |
type base_slice = int * bool * bool * int * string |
75026 | 16 |
type atp_slice = atp_format * string * string * bool * string |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
17 |
type atp_config = |
73374 | 18 |
{exec : string list * string list, |
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77269
diff
changeset
|
19 |
arguments : bool -> bool -> string -> Time.time -> Path.T -> string list, |
42578
1eaf4d437d4c
define type system in ATP module so that ATPs can suggest a type system
blanchet
parents:
42577
diff
changeset
|
20 |
proof_delims : (string * string) list, |
53586
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
53515
diff
changeset
|
21 |
known_failures : (atp_failure * string) list, |
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
53515
diff
changeset
|
22 |
prem_role : atp_formula_role, |
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
23 |
generate_isabelle_info : bool, |
75026 | 24 |
good_slices : Proof.context -> (base_slice * atp_slice) list, |
25 |
good_max_mono_iters : int, |
|
26 |
good_max_new_mono_instances : int} |
|
38023 | 27 |
|
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
28 |
val default_max_mono_iters : int |
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
29 |
val default_max_new_mono_instances : int |
50950 | 30 |
val spass_H1SOS : string |
31 |
val spass_H2 : string |
|
32 |
val spass_H2LR0LT0 : string |
|
33 |
val spass_H2NuVS0 : string |
|
34 |
val spass_H2NuVS0Red2 : string |
|
35 |
val spass_H2SOS : string |
|
73435
1cc848548f21
invoke remote ATP via SystemOnTPTP.run_systems from Isabelle/Scala (without perl);
wenzelm
parents:
73432
diff
changeset
|
36 |
val isabelle_scala_function: string list * string list |
57671
dc5e1b1db9ba
avoid 'eproof' and 'eproof_ram' scripts if possible (i.e. if 'eprover' can produce reasonable enough proofs for one-liner reconstruction)
blanchet
parents:
57636
diff
changeset
|
37 |
val remote_atp : string -> string -> string list -> (string * string) list -> |
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
38 |
(atp_failure * string) list -> atp_formula_role -> bool -> |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
39 |
(Proof.context -> base_slice * atp_slice) -> string * (unit -> atp_config) |
47606
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents:
47506
diff
changeset
|
40 |
val add_atp : string * (unit -> atp_config) -> theory -> theory |
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents:
47506
diff
changeset
|
41 |
val get_atp : theory -> string -> (unit -> atp_config) |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
42 |
val is_atp_installed : theory -> string -> bool |
35867 | 43 |
val refresh_systems_on_tptp : unit -> unit |
75036 | 44 |
val local_atps : string list |
45 |
val remote_atps : string list |
|
75465
d9b23902692d
excluded dummy ATPs from Sledgehammer's default provers
desharna
parents:
75433
diff
changeset
|
46 |
val dummy_atps : string list |
d9b23902692d
excluded dummy ATPs from Sledgehammer's default provers
desharna
parents:
75433
diff
changeset
|
47 |
val non_dummy_atps : string list |
d9b23902692d
excluded dummy ATPs from Sledgehammer's default provers
desharna
parents:
75433
diff
changeset
|
48 |
val all_atps : string list |
28592 | 49 |
end; |
50 |
||
72400 | 51 |
structure Sledgehammer_ATP_Systems : SLEDGEHAMMER_ATP_SYSTEMS = |
28592 | 52 |
struct |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
53 |
|
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
|
54 |
open ATP_Problem |
39491
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents:
39375
diff
changeset
|
55 |
open ATP_Proof |
46320 | 56 |
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
|
57 |
|
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
58 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
59 |
(* ATP configuration *) |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
60 |
|
74890 | 61 |
val TF0 = TFF (Monomorphic, Without_FOOL) |
62 |
val TF1 = TFF (Polymorphic, Without_FOOL) |
|
63 |
val TX0 = TFF (Monomorphic, With_FOOL {with_ite = true, with_let = true}) |
|
79796 | 64 |
val TX0_MINUS = TFF (Monomorphic, With_FOOL {with_ite = false, with_let = false}) |
74890 | 65 |
val TX1 = TFF (Polymorphic, With_FOOL {with_ite = true, with_let = true}) |
66 |
val TH0 = THF (Monomorphic, {with_ite = true, with_let = true}, THF_With_Choice) |
|
78788 | 67 |
val TH0_MINUS = THF (Monomorphic, {with_ite = false, with_let = false}, THF_Without_Choice) |
74895
b605ebd87a47
proper polymorphism for TH1 format in Sledgehammer
desharna
parents:
74894
diff
changeset
|
68 |
val TH1 = THF (Polymorphic, {with_ite = true, with_let = true}, THF_With_Choice) |
78788 | 69 |
val TH1_MINUS = THF (Polymorphic, {with_ite = false, with_let = false}, THF_Without_Choice) |
74890 | 70 |
|
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
71 |
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
|
72 |
val default_max_new_mono_instances = 100 (* FUDGE *) |
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
73 |
|
77428 | 74 |
(* desired slice size, abduce, falsify, desired number of facts, fact filter *) |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
76301
diff
changeset
|
75 |
type base_slice = int * bool * bool * int * string |
75339
d9bb81999d2c
first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
blanchet
parents:
75059
diff
changeset
|
76 |
|
75340
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
77 |
(* problem file format, type encoding, lambda translation scheme, uncurried aliases?, |
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
78 |
prover-specific extra information *) |
75026 | 79 |
type atp_slice = atp_format * string * string * bool * string |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46407
diff
changeset
|
80 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
81 |
type atp_config = |
73374 | 82 |
{exec : string list * string list, |
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77269
diff
changeset
|
83 |
arguments : bool -> bool -> string -> Time.time -> Path.T -> string list, |
42578
1eaf4d437d4c
define type system in ATP module so that ATPs can suggest a type system
blanchet
parents:
42577
diff
changeset
|
84 |
proof_delims : (string * string) list, |
53586
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
53515
diff
changeset
|
85 |
known_failures : (atp_failure * string) list, |
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
53515
diff
changeset
|
86 |
prem_role : atp_formula_role, |
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
87 |
generate_isabelle_info : bool, |
75026 | 88 |
good_slices : Proof.context -> (base_slice * atp_slice) list, |
89 |
good_max_mono_iters : int, |
|
90 |
good_max_new_mono_instances : int} |
|
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
91 |
|
75339
d9bb81999d2c
first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
blanchet
parents:
75059
diff
changeset
|
92 |
(* "good_slices" must be found empirically, ideally taking a holistic approach since the ATPs are |
d9bb81999d2c
first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
blanchet
parents:
75059
diff
changeset
|
93 |
run in parallel. *) |
42710
84fcce345b5d
further improved type system setup based on Judgment Days
blanchet
parents:
42709
diff
changeset
|
94 |
|
51011 | 95 |
val mepoN = "mepo" |
96 |
val mashN = "mash" |
|
97 |
val meshN = "mesh" |
|
98 |
||
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
99 |
val tstp_proof_delims = |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
100 |
[("% SZS output start CNFRefutation", "% SZS output end CNFRefutation"), |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
101 |
("% SZS output start Refutation", "% SZS output end Refutation"), |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
102 |
("% SZS output start Proof", "% SZS output end Proof")] |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
103 |
|
45203 | 104 |
fun known_szs_failures wrap = |
105 |
[(Unprovable, wrap "CounterSatisfiable"), |
|
106 |
(Unprovable, wrap "Satisfiable"), |
|
107 |
(GaveUp, wrap "GaveUp"), |
|
108 |
(GaveUp, wrap "Unknown"), |
|
109 |
(GaveUp, wrap "Incomplete"), |
|
110 |
(ProofMissing, wrap "Theorem"), |
|
111 |
(ProofMissing, wrap "Unsatisfiable"), |
|
112 |
(TimedOut, wrap "Timeout"), |
|
113 |
(Inappropriate, wrap "Inappropriate"), |
|
114 |
(OutOfResources, wrap "ResourceOut"), |
|
115 |
(OutOfResources, wrap "MemoryOut"), |
|
116 |
(Interrupted, wrap "Forced"), |
|
117 |
(Interrupted, wrap "User")] |
|
118 |
||
119 |
val known_szs_status_failures = known_szs_failures (prefix "SZS status ") |
|
120 |
val known_says_failures = known_szs_failures (prefix " says ") |
|
121 |
||
38023 | 122 |
structure Data = Theory_Data |
123 |
( |
|
47606
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents:
47506
diff
changeset
|
124 |
type T = ((unit -> atp_config) * stamp) Symtab.table |
38023 | 125 |
val empty = Symtab.empty |
46407
30e9720cc0b9
optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents:
46402
diff
changeset
|
126 |
fun merge data : T = |
30e9720cc0b9
optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents:
46402
diff
changeset
|
127 |
Symtab.merge (eq_snd (op =)) data |
63692 | 128 |
handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name) |
38023 | 129 |
) |
38017
3ad3e3ca2451
move Sledgehammer-specific code out of "Sledgehammer_TPTP_Format"
blanchet
parents:
38015
diff
changeset
|
130 |
|
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
|
131 |
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
|
132 |
|
43473
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
blanchet
parents:
43467
diff
changeset
|
133 |
val sosN = "sos" |
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
blanchet
parents:
43467
diff
changeset
|
134 |
val no_sosN = "no_sos" |
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
blanchet
parents:
43467
diff
changeset
|
135 |
|
47032 | 136 |
|
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
137 |
(* agsyHOL *) |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
138 |
|
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
139 |
val agsyhol_config : atp_config = |
73374 | 140 |
{exec = (["AGSYHOL_HOME"], ["agsyHOL"]), |
75341 | 141 |
arguments = fn _ => fn _ => fn _ => fn timeout => fn problem => |
73568
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
142 |
["--proof --time-out " ^ string_of_int (to_secs 1 timeout) ^ " " ^ File.bash_path problem], |
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
143 |
proof_delims = tstp_proof_delims, |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
144 |
known_failures = known_szs_status_failures, |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
145 |
prem_role = Hypothesis, |
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
146 |
generate_isabelle_info = false, |
75026 | 147 |
good_slices = |
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
148 |
(* FUDGE *) |
82231
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
149 |
K [((2, false, false, 60, meshN), (THF (Monomorphic, {with_ite = false, with_let = false}, THF_Without_Choice), "mono_native_higher", keep_lamsN, false, ""))], |
75026 | 150 |
good_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *), |
151 |
good_max_new_mono_instances = default_max_new_mono_instances} |
|
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
152 |
|
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
153 |
val agsyhol = (agsyholN, fn () => agsyhol_config) |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
154 |
|
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
155 |
|
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
156 |
(* Alt-Ergo *) |
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
157 |
|
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
158 |
val alt_ergo_config : atp_config = |
73374 | 159 |
{exec = (["WHY3_HOME"], ["why3"]), |
75341 | 160 |
arguments = fn _ => fn _ => fn _ => fn timeout => fn problem => |
73568
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
161 |
["--format tptp --prover 'Alt-Ergo,0.95.2,' --timelimit " ^ string_of_int (to_secs 1 timeout) ^ |
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
162 |
" " ^ File.bash_path problem], |
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
163 |
proof_delims = [], |
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
164 |
known_failures = |
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
165 |
[(ProofMissing, ": Valid"), |
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
166 |
(TimedOut, ": Timeout"), |
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
167 |
(GaveUp, ": Unknown")], |
47976 | 168 |
prem_role = Hypothesis, |
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
169 |
generate_isabelle_info = false, |
75026 | 170 |
good_slices = |
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
171 |
(* FUDGE *) |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
76301
diff
changeset
|
172 |
K [((1000 (* infinity *), false, false, 100, meshN), (TF1, "poly_native", liftingN, false, ""))], |
75026 | 173 |
good_max_mono_iters = default_max_mono_iters, |
174 |
good_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
|
175 |
|
47646 | 176 |
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
|
177 |
|
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
178 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
179 |
(* E *) |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
180 |
|
79796 | 181 |
local |
182 |
||
183 |
fun make_e_config max_new_mono_insts good_slices : atp_config = |
|
73973 | 184 |
{exec = (["E_HOME"], ["eprover-ho", "eprover"]), |
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77269
diff
changeset
|
185 |
arguments = fn abduce => fn _ => fn extra_options => fn timeout => fn problem => |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77269
diff
changeset
|
186 |
["--tstp-in --tstp-out --silent " ^ |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77269
diff
changeset
|
187 |
(if abduce then |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77269
diff
changeset
|
188 |
"--auto --print-saturated=ieIE --print-sat-info --soft-cpu-limit=" |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77269
diff
changeset
|
189 |
else |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77269
diff
changeset
|
190 |
extra_options ^ " --cpu-limit=") ^ |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77269
diff
changeset
|
191 |
string_of_int (to_secs 2 timeout) ^ " --proof-object=1 " ^ File.bash_path problem], |
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
192 |
proof_delims = |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
193 |
[("# SZS output start CNFRefutation", "# SZS output end CNFRefutation")] @ |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
194 |
tstp_proof_delims, |
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
195 |
known_failures = |
45203 | 196 |
[(TimedOut, "Failure: Resource limit exceeded (time)"), |
47972 | 197 |
(TimedOut, "time limit exceeded")] @ |
198 |
known_szs_status_failures, |
|
47976 | 199 |
prem_role = Conjecture, |
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
200 |
generate_isabelle_info = false, |
79796 | 201 |
good_slices = K good_slices, |
202 |
good_max_mono_iters = default_max_mono_iters, |
|
203 |
good_max_new_mono_instances = max_new_mono_insts} |
|
204 |
||
82382 | 205 |
val e_config : atp_config = |
79796 | 206 |
(* FUDGE: this solved 950/1500 (63.33 %) using MePo when testing *) |
82200 | 207 |
let |
208 |
val extra_options = "--auto-schedule" |
|
209 |
val good_slices = |
|
82231
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
210 |
[(((2, false, false, 128, meshN), (TX0_MINUS, "mono_native_fool", liftingN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
211 |
(((2, false, false, 1024, meshN), (TH0, "mono_native_higher", keep_lamsN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
212 |
(((2, false, false, 128, meshN), (TH0_MINUS, "mono_native_higher", keep_lamsN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
213 |
(((2, false, false, 2048, meshN), (TF0, "mono_native", combsN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
214 |
(((2, false, false, 512, meshN), (TF0, "mono_native", liftingN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
215 |
(((2, false, false, 1024, meshN), (TX0_MINUS, "mono_native_fool", liftingN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
216 |
(((2, false, false, 64, meshN), (TH0, "mono_native_higher", keep_lamsN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
217 |
(((2, false, false, 512, meshN), (TH0_MINUS, "mono_native_higher", keep_lamsN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
218 |
(((2, false, false, 32, meshN), (TF0, "mono_native", liftingN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
219 |
(((2, false, false, 2048, meshN), (TH0, "mono_native_higher", keep_lamsN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
220 |
(((2, false, false, 256, meshN), (TH0, "mono_native_higher", keep_lamsN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
221 |
(((2, false, false, 512, meshN), (TF0, "mono_native", combsN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
222 |
(((2, false, false, 1024, meshN), (TF0, "mono_native", liftingN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
223 |
(((2, false, false, 16, meshN), (TH0, "mono_native_higher", keep_lamsN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
224 |
(((2, false, false, 512, meshN), (TX0_MINUS, "mono_native_fool", liftingN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
225 |
(((2, false, false, 64, meshN), (TF0, "mono_native", liftingN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
226 |
(((2, false, false, 128, meshN), (TF0, "mono_native", combsN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
227 |
(((2, false, false, 2048, meshN), (TH0_MINUS, "mono_native_higher", keep_lamsN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
228 |
(((2, false, false, 128, meshN), (TX0_MINUS, "mono_native_fool", combsN, false, extra_options))), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
229 |
(((2, false, false, 2048, meshN), (TX0_MINUS, "mono_native_fool", liftingN, false, extra_options)))] |
82200 | 230 |
in |
231 |
make_e_config 128 good_slices |
|
232 |
end |
|
79796 | 233 |
|
234 |
in |
|
235 |
||
82382 | 236 |
val e = (eN, fn () => e_config) |
79796 | 237 |
|
238 |
end |
|
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
239 |
|
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
240 |
|
48700 | 241 |
(* iProver *) |
242 |
||
243 |
val iprover_config : atp_config = |
|
73374 | 244 |
{exec = (["IPROVER_HOME"], ["iproveropt", "iprover"]), |
75341 | 245 |
arguments = fn _ => fn _ => fn _ => fn timeout => fn problem => |
74350 | 246 |
["--clausifier \"$VAMPIRE_HOME\"/vampire " ^ |
74046
462d652ad910
use Vampire's clausifier with iProver, now that E's is no longer supported
blanchet
parents:
74005
diff
changeset
|
247 |
"--clausifier_options \"--mode clausify\" " ^ |
73568
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
248 |
"--time_out_real " ^ string_of_real (Time.toReal timeout) ^ " " ^ File.bash_path problem], |
48700 | 249 |
proof_delims = tstp_proof_delims, |
250 |
known_failures = |
|
251 |
[(ProofIncomplete, "% SZS output start CNFRefutation")] @ |
|
252 |
known_szs_status_failures, |
|
253 |
prem_role = Hypothesis, |
|
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
254 |
generate_isabelle_info = false, |
75026 | 255 |
good_slices = |
48700 | 256 |
(* FUDGE *) |
82231
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
257 |
K [((2, false, false, 32, meshN), (TF0, "mono_native", liftingN, false, "")), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
258 |
((2, false, false, 512, meshN), (TX0, "mono_native", liftingN, false, "")), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
259 |
((2, false, false, 128, mashN), (TF0, "mono_native", combsN, false, "")), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
260 |
((2, false, false, 1024, meshN), (TF0, "mono_native", liftingN, false, "")), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
261 |
((2, false, true, 256, mepoN), (TF0, "mono_native", combsN, false, ""))], |
75026 | 262 |
good_max_mono_iters = default_max_mono_iters, |
263 |
good_max_new_mono_instances = default_max_new_mono_instances} |
|
48700 | 264 |
|
265 |
val iprover = (iproverN, fn () => iprover_config) |
|
266 |
||
267 |
||
44099 | 268 |
(* LEO-II *) |
269 |
||
270 |
val leo2_config : atp_config = |
|
73374 | 271 |
{exec = (["LEO2_HOME"], ["leo.opt", "leo"]), |
75341 | 272 |
arguments = fn _ => fn full_proofs => fn _ => fn timeout => fn problem => |
73568
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
273 |
["--foatp e --atp e=\"$E_HOME\"/eprover \ |
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
274 |
\--atp epclextract=\"$E_HOME\"/epclextract \ |
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
275 |
\--proofoutput 1 --timeout " ^ string_of_int (to_secs 1 timeout) ^ " " ^ |
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
276 |
(if full_proofs then "--notReplLeibnizEQ --notReplAndrewsEQ --notUseExtCnfCmbd " else "") ^ |
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
277 |
File.bash_path problem], |
44099 | 278 |
proof_delims = tstp_proof_delims, |
45207 | 279 |
known_failures = |
47974
08d2dcc2dab9
improved LEO-II definition handling -- still hoping for a fix directly in LEO-II
blanchet
parents:
47972
diff
changeset
|
280 |
[(TimedOut, "CPU time limit exceeded, terminating"), |
47972 | 281 |
(GaveUp, "No.of.Axioms")] @ |
282 |
known_szs_status_failures, |
|
47976 | 283 |
prem_role = Hypothesis, |
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
284 |
generate_isabelle_info = false, |
75026 | 285 |
good_slices = |
44099 | 286 |
(* FUDGE *) |
82231
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
287 |
K [((2, false, false, 40, meshN), (THF (Monomorphic, {with_ite = false, with_let = false}, THF_Without_Choice), "mono_native_higher", keep_lamsN, false, ""))], |
75026 | 288 |
good_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *), |
289 |
good_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
|
290 |
|
47646 | 291 |
val leo2 = (leo2N, fn () => leo2_config) |
44099 | 292 |
|
293 |
||
67021
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66544
diff
changeset
|
294 |
(* Leo-III *) |
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66544
diff
changeset
|
295 |
|
69717 | 296 |
(* Include choice? Disabled now since it's disabled for Satallax as well. *) |
67021
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66544
diff
changeset
|
297 |
val leo3_config : atp_config = |
73374 | 298 |
{exec = (["LEO3_HOME"], ["leo3"]), |
75341 | 299 |
arguments = fn _ => fn full_proofs => fn _ => fn timeout => fn problem => |
82024
bbda3b4f3c99
switch from CVC5 to cvc5, including updates of internal tool references;
wenzelm
parents:
81741
diff
changeset
|
300 |
[File.bash_path problem ^ " " ^ "--atp cvc=\"$CVC5_SOLVER\" --atp e=\"$E_HOME\"/eprover \ |
73568
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
301 |
\-p -t " ^ string_of_int (to_secs 1 timeout) ^ " " ^ |
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
302 |
(if full_proofs then "--nleq --naeq " else "")], |
67021
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66544
diff
changeset
|
303 |
proof_delims = tstp_proof_delims, |
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66544
diff
changeset
|
304 |
known_failures = known_szs_status_failures, |
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66544
diff
changeset
|
305 |
prem_role = Hypothesis, |
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
306 |
generate_isabelle_info = false, |
75026 | 307 |
good_slices = |
67021
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66544
diff
changeset
|
308 |
(* FUDGE *) |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
76301
diff
changeset
|
309 |
K [((6, false, false, 512, meshN), (TH0, "mono_native_higher", keep_lamsN, false, "")), |
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
76301
diff
changeset
|
310 |
((6, false, false, 512, meshN), (TF0, "mono_native", liftingN, false, ""))], |
75026 | 311 |
good_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *), |
312 |
good_max_new_mono_instances = default_max_new_mono_instances} |
|
67021
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66544
diff
changeset
|
313 |
|
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66544
diff
changeset
|
314 |
val leo3 = (leo3N, fn () => leo3_config) |
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66544
diff
changeset
|
315 |
|
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66544
diff
changeset
|
316 |
|
44099 | 317 |
(* Satallax *) |
318 |
||
52097 | 319 |
(* Choice is disabled until there is proper reconstruction for it. *) |
44099 | 320 |
val satallax_config : atp_config = |
73374 | 321 |
{exec = (["SATALLAX_HOME"], ["satallax.opt", "satallax"]), |
75341 | 322 |
arguments = fn _ => fn _ => fn _ => fn timeout => fn problem => |
73568
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
323 |
[(case getenv "E_HOME" of |
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
324 |
"" => "" |
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
325 |
| home => "-E " ^ home ^ "/eprover ") ^ |
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
326 |
"-p tstp -t " ^ string_of_int (to_secs 1 timeout) ^ " " ^ File.bash_path problem], |
45162 | 327 |
proof_delims = |
57707
0242e9578828
imported patch satallax_proof_support_Sledgehammer
fleury
parents:
57672
diff
changeset
|
328 |
[("% SZS output start Proof", "% SZS output end Proof")], |
45203 | 329 |
known_failures = known_szs_status_failures, |
47981 | 330 |
prem_role = Hypothesis, |
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
331 |
generate_isabelle_info = false, |
75026 | 332 |
good_slices = |
44754 | 333 |
(* FUDGE *) |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
76301
diff
changeset
|
334 |
K [((12, false, false, 256, meshN), (THF (Monomorphic, {with_ite = false, with_let = false}, THF_Without_Choice), "mono_native_higher", keep_lamsN, false, ""))], |
75026 | 335 |
good_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *), |
336 |
good_max_new_mono_instances = default_max_new_mono_instances} |
|
44099 | 337 |
|
47646 | 338 |
val satallax = (satallaxN, fn () => satallax_config) |
44099 | 339 |
|
340 |
||
341 |
(* 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
|
342 |
|
48005
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents:
48004
diff
changeset
|
343 |
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
|
344 |
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
|
345 |
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
|
346 |
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
|
347 |
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
|
348 |
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
|
349 |
|
48005
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents:
48004
diff
changeset
|
350 |
val spass_config : atp_config = |
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
351 |
{exec = (["SPASS_HOME"], ["SPASS"]), |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
352 |
arguments = fn _ => fn full_proofs => fn extra_options => fn timeout => fn problem => |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
353 |
["-Isabelle=1 " ^ (if full_proofs then "-CNFRenaming=0 -Splits=0 " else "") ^ |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
354 |
"-TimeLimit=" ^ string_of_int (to_secs 1 timeout) ^ " " ^ File.bash_path problem |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
355 |
|> extra_options <> "" ? prefix (extra_options ^ " ")], |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
356 |
proof_delims = [("Here is a proof", "Formulae used in the proof")], |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
357 |
known_failures = |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
358 |
[(GaveUp, "SPASS beiseite: Completion found"), |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
359 |
(TimedOut, "SPASS beiseite: Ran out of time"), |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
360 |
(OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"), |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
361 |
(MalformedInput, "Undefined symbol"), |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
362 |
(MalformedInput, "Free Variable"), |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
363 |
(Unprovable, "No formulae and clauses found in input file"), |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
364 |
(InternalError, "Please report this error")], |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
365 |
prem_role = Conjecture, |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
366 |
generate_isabelle_info = true, |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
367 |
good_slices = |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
368 |
(* FUDGE *) |
82231
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
369 |
K [((2, false, false, 150, meshN), (DFG Monomorphic, "mono_native", combsN, true, "")), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
370 |
((2, false, false, 500, meshN), (DFG Monomorphic, "mono_native", liftingN, true, spass_H2SOS)), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
371 |
((2, false, false, 50, meshN), (DFG Monomorphic, "mono_native", liftingN, true, spass_H2LR0LT0)), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
372 |
((2, false, false, 250, meshN), (DFG Monomorphic, "mono_native", combsN, true, spass_H2NuVS0)), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
373 |
((2, false, false, 1000, mepoN), (DFG Monomorphic, "mono_native", liftingN, true, spass_H1SOS)), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
374 |
((2, false, true, 150, meshN), (DFG Monomorphic, "poly_guards??", liftingN, false, spass_H2NuVS0Red2)), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
375 |
((2, false, false, 300, meshN), (DFG Monomorphic, "mono_native", combsN, true, spass_H2SOS)), |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
376 |
((2, false, false, 100, meshN), (DFG Monomorphic, "mono_native", combs_and_liftingN, true, spass_H2))], |
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
377 |
good_max_mono_iters = default_max_mono_iters, |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
378 |
good_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
|
379 |
|
48005
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents:
48004
diff
changeset
|
380 |
val spass = (spassN, fn () => spass_config) |
38454
9043eefe8d71
detect old Vampire and give a nicer error message
blanchet
parents:
38433
diff
changeset
|
381 |
|
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
382 |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
383 |
(* Vampire *) |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
384 |
|
78788 | 385 |
local |
386 |
||
387 |
val new_vampire_basic_options = |
|
388 |
["--input_syntax tptp", |
|
389 |
"--proof tptp", |
|
390 |
"--output_axiom_names on"] @ |
|
391 |
(if ML_System.platform_is_windows |
|
392 |
then [] (*time slicing is not support in the Windows version of Vampire*) |
|
393 |
else ["--mode portfolio", "--schedule snake_slh"]) |
|
58084 | 394 |
|
395 |
val vampire_full_proof_options = |
|
78788 | 396 |
["--proof_extra free", |
397 |
"--forced_options avatar=off:equality_proxy=off:general_splitting=off:inequality_splitting=0:naming=0"] |
|
58084 | 398 |
|
78788 | 399 |
fun make_vampire_config vampire_basic_options good_max_new_mono_instances good_slices : atp_config = |
74890 | 400 |
{exec = (["VAMPIRE_HOME"], ["vampire"]), |
78788 | 401 |
arguments = fn _ => fn full_proofs => fn extra_options => fn timeout => fn problem => |
402 |
vampire_basic_options @ |
|
403 |
(if full_proofs then vampire_full_proof_options else []) @ |
|
404 |
(if extra_options <> "" then [extra_options] else []) @ |
|
405 |
["-t " ^ string_of_int (to_secs 1 timeout), |
|
406 |
"--input_file " ^ File.bash_path problem], |
|
74890 | 407 |
proof_delims = |
408 |
[("=========== Refutation ==========", |
|
409 |
"======= End of refutation =======")] @ |
|
410 |
tstp_proof_delims, |
|
411 |
known_failures = |
|
412 |
[(GaveUp, "UNPROVABLE"), |
|
413 |
(GaveUp, "CANNOT PROVE"), |
|
414 |
(Unprovable, "Satisfiability detected"), |
|
415 |
(Unprovable, "Termination reason: Satisfiable"), |
|
416 |
(Interrupted, "Aborted by signal SIGINT")] @ |
|
417 |
known_szs_status_failures, |
|
418 |
prem_role = Hypothesis, |
|
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
419 |
generate_isabelle_info = false, |
78788 | 420 |
good_slices = K good_slices, |
75026 | 421 |
good_max_mono_iters = default_max_mono_iters, |
78788 | 422 |
good_max_new_mono_instances = good_max_new_mono_instances} |
423 |
||
82383
7ed32d7f8317
removed old Vampire configuration from Sledgehammer
desharna
parents:
82382
diff
changeset
|
424 |
val vampire_config : atp_config = |
79797 | 425 |
(* FUDGE: this solved 998/1500 (66.53 %) using MePo when testing *) |
426 |
make_vampire_config new_vampire_basic_options 256 |
|
78789
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
427 |
[(((2, false, false, 512, meshN), (TX0, "mono_native_fool", liftingN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
428 |
(((2, false, false, 2048, meshN), (TX0, "mono_native_fool", combsN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
429 |
(((2, false, false, 128, meshN), (TX0, "mono_native_fool", liftingN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
430 |
(((2, false, false, 1024, meshN), (TF1, "poly_native", liftingN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
431 |
(((2, false, false, 256, meshN), (TH0_MINUS, "mono_native_higher", keep_lamsN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
432 |
(((2, false, false, 1024, meshN), (TX0, "mono_native_fool", liftingN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
433 |
(((2, false, false, 256, meshN), (TX0, "mono_native_fool", liftingN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
434 |
(((2, false, false, 2048, meshN), (TF1, "poly_native", combsN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
435 |
(((2, false, false, 256, meshN), (TX1, "poly_native_fool", liftingN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
436 |
(((2, false, false, 512, meshN), (TF0, "mono_native", liftingN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
437 |
(((2, false, false, 2048, meshN), (TF0, "mono_native", liftingN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
438 |
(((2, false, false, 64, meshN), (TF1, "poly_native", combsN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
439 |
(((2, false, false, 256, meshN), (TH1_MINUS, "poly_native_higher", keep_lamsN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
440 |
(((2, false, false, 256, meshN), (TF0, "mono_native", combsN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
441 |
(((2, false, false, 256, meshN), (TF0, "mono_native", liftingN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
442 |
(((2, false, false, 32, meshN), (TH0_MINUS, "mono_native_higher", keep_lamsN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
443 |
(((2, false, false, 512, meshN), (TH0_MINUS, "mono_native_higher", keep_lamsN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
444 |
(((2, false, false, 512, meshN), (TX0, "mono_native_fool", combsN, false, ""))), |
f2e845c3e65c
proper slice duration (i.e., 5 s) for new Vampire portfolio following 5a14f2cc1ea0
desharna
parents:
78788
diff
changeset
|
445 |
(((2, false, false, 1024, meshN), (TX1, "poly_native_fool", combsN, false, "")))] |
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
446 |
|
78788 | 447 |
in |
448 |
||
82383
7ed32d7f8317
removed old Vampire configuration from Sledgehammer
desharna
parents:
82382
diff
changeset
|
449 |
val vampire = (vampireN, fn () => vampire_config) |
78788 | 450 |
|
451 |
end |
|
68563 | 452 |
|
69717 | 453 |
(* Zipperposition *) |
454 |
||
57154 | 455 |
val zipperposition_config : atp_config = |
73375 | 456 |
let |
74890 | 457 |
val format = |
458 |
THF (Polymorphic, {with_ite = true, with_let = false}, THF_Without_Choice) |
|
73375 | 459 |
in |
460 |
{exec = (["ZIPPERPOSITION_HOME"], ["zipperposition"]), |
|
75341 | 461 |
arguments = fn _ => fn _ => fn extra_options => fn timeout => fn problem => |
74471 | 462 |
["--input tptp", "--output tptp", "--timeout " ^ Time.toString timeout, extra_options, |
463 |
File.bash_path problem], |
|
73375 | 464 |
proof_delims = tstp_proof_delims, |
74207 | 465 |
known_failures = |
466 |
[(TimedOut, "SZS status ResourceOut")] @ (* odd way of timing out *) |
|
467 |
known_szs_status_failures, |
|
73375 | 468 |
prem_role = Hypothesis, |
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
469 |
generate_isabelle_info = true, |
75026 | 470 |
good_slices = |
82231
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
471 |
K [((1, false, false, 512, meshN), (format, "mono_native_higher_fool", keep_lamsN, false, "--mode=ho-pragmatic --tptp-def-as-rewrite --rewrite-before-cnf=true --max-inferences=1 --ho-unif-max-depth=1 --ho-max-elims=0 --ho-max-app-projections=0 --ho-max-rigid-imitations=1 --ho-max-identifications=0 --boolean-reasoning=bool-hoist --bool-hoist-simpl=true --bool-select=LI --recognize-injectivity=true --ext-rules=ext-family --ext-rules-max-depth=1 --ho-choice-inst=true --ho-prim-enum=none --ho-elim-leibniz=0 --interpret-bool-funs=true --try-e=\"$E_HOME/eprover\" --tmp-dir=\"$ISABELLE_TMP_PREFIX\" --ho-unif-level=pragmatic-framework --select=bb+e-selection2 --post-cnf-lambda-lifting=true -q \"4|prefer-sos|pnrefined(2,1,1,1,2,2,2)\" -q \"6|prefer-processed|conjecture-relative-struct(1.5,3.5,2,3)\" -q \"1|const|fifo\" -q \"4|prefer-ground|orient-lmax(2,1,2,1,1)\" -q \"4|defer-sos|conjecture-relative-struct(1,5,2,3)\" --avatar=off --recognize-injectivity=true --ho-neg-ext=true --e-timeout=2 --ho-pattern-decider=true --ho-fixpoint-decider=true --e-max-derived=50 --ignore-orphans=true --e-auto=true --presaturate=true --e-call-point=0.1")), (* sh5_sh1.sh *) |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
472 |
((1, false, false, 512, mepoN), (format, "mono_native_higher_fool", keep_lamsN, false, "--mode=ho-pragmatic -nc --tptp-def-as-rewrite --rewrite-before-cnf=true --mode=ho-competitive --boolean-reasoning=simpl-only --ext-rules=ext-family --ext-rules-max-depth=1 --ho-prim-enum=none --avatar=off --recognize-injectivity=true --ho-elim-leibniz=1 --ho-unif-level=pragmatic-framework --no-max-vars --max-inferences=4 --ho-max-app-projections=1 --ho-max-elims=0 --ho-max-rigid-imitations=2 --ho-max-identifications=0 --ho-unif-max-depth=3 -q \"6|prefer-sos|pnrefined(1,1,1,2,2,2,0.5)\" -q \"6|const|conjecture-relative-var(1.02,l,f)\" -q \"1|prefer-processed|fifo\" -q \"1|prefer-non-goals|conjecture-relative-var(1,l,f)\" -q \"4|prefer-easy-ho|conjecture-relative-var(1.01,s,f)\" --select=e-selection7 --ho-choice-inst=true --try-e=\"$E_HOME/eprover\" --tmp-dir=\"$ISABELLE_TMP_PREFIX\" --e-timeout=7 --sine=50 --sine-tolerance=1 --sine-depth-max=2 --sine-depth-min=1 --e-max-derived=64 --sine-ignore-k-most-common-syms=2 --sine-trim-implications=true --e-encode-lambdas=lift --scan-clause-ac=false --lambdasup=0 --kbo-weight-fun=lambda-def-invfreqrank --demod-in-var-args=true --bool-demod=true --lambda-demod=true --e-call-point=0.1")), (* sh10_c_ic.sh *) |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
473 |
((1, false, false, 1024, meshN), (format, "mono_native_higher_fool", keep_lamsN, false, "--mode=ho-pragmatic --tptp-def-as-rewrite --rewrite-before-cnf=true --mode=ho-competitive --boolean-reasoning=simpl-only --ext-rules=ext-family --ext-rules-max-depth=1 --ho-prim-enum=none --avatar=off --recognize-injectivity=true --ho-elim-leibniz=1 --ho-unif-level=pragmatic-framework --no-max-vars --max-inferences=2 --ho-unif-max-depth=1 -q \"6|prefer-sos|pnrefined(1,1,1,2,2,2,0.5)\" -q \"6|const|conjecture-relative-var(1.02,l,f)\" -q \"1|prefer-processed|fifo\" -q \"1|prefer-non-goals|conjecture-relative-var(1,l,f)\" -q \"4|prefer-easy-ho|conjecture-relative-var(1.01,s,f)\" --select=e-selection16 --ho-choice-inst=true --try-e=\"$E_HOME/eprover\" --tmp-dir=\"$ISABELLE_TMP_PREFIX\" --e-timeout=3 --e-auto=true --sine=50 --sine-tolerance=1.0 --sine-depth-max=3 --sine-depth-min=1 --sine-trim-implications=true --ho-unif-level=pragmatic-framework --e-encode-lambdas=lift --scan-clause-ac=false --kbo-weight-fun=lambda-def-invfreqrank --e-call-point=0.1")), (* sh8_shallow_sine.sh *) |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
474 |
((1, false, false, 256, mepoN), (format, "mono_native_higher_fool", keep_lamsN, false, "-nc --tptp-def-as-rewrite --rewrite-before-cnf=true --mode=ho-competitive --boolean-reasoning=simpl-only --ext-rules=off --ho-prim-enum=full --ho-prim-max=1 --avatar=off --recognize-injectivity=true --ho-elim-leibniz=4 --ho-unif-level=full-framework --no-max-vars -q \"2|prefer-goals|conjecture-relative-e(0.5,1,100,100,100,100,1.5,1.5,1)\" -q \"4|const|conjecture-relative-e(0.1,1,100,100,100,100,1.5,1.5,1.5)\" -q \"1|prefer-processed|fifo\" -q \"1|prefer-non-goals|conjecture-relative-e(0.5,1,100,100,100,100,1.5,1.5,1.5)\" -q \"4|prefer-sos|pnrefined(1,1,1,1,2,1.5,2)\" --select=ho-selection5 --ho-choice-inst=true --try-e=\"$E_HOME/eprover\" --tmp-dir=\"$ISABELLE_TMP_PREFIX\" --e-timeout=5 --e-call-point=0.25 --e-auto=true --sine=50 --sine-tolerance=2 --sine-depth-max=4 --sine-depth-min=1 --e-max-derived=96 --e-encode-lambdas=lift --scan-clause-ac=false --kbo-weight-fun=arity0 --prec-gen-fun=invfreq_conj")), (* sh10_new_c.s3.sh *) |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
475 |
((1, false, true, 512, meshN), (format, "mono_native_higher_fool", keep_lamsN, false, "--mode=ho-comb-complete --boolean-reasoning=simpl-only --ext-rules=off --kbo-weight-fun=lambda-def-sqarity --ho-prim-enum=none --tptp-def-as-rewrite -q \"4|prefer-sos|orient-lmax(2,1,2,1,1)\" -q \"4|defer-sos|conjecture-relative-var(1,s,f)\" -q \"3|const|default\" -q \"1|prefer-processed|fifo\" --ho-elim-leibniz=1 --select=NoSelection --solve-formulas=true --lazy-cnf=true --lazy-cnf-kind=simp --lazy-cnf-renaming-threshold=8 --sine=60 --sine-tolerance=2 --sine-depth-max=5 --sine-depth-min=1 --try-e=\"$E_HOME/eprover\" --tmp-dir=\"$ISABELLE_TMP_PREFIX\" --e-timeout=3 --e-auto=true --e-max-derived=50 --e-encode-lambdas=lift --scan-clause-ac=false --presaturate=true --comb-b-penalty=3 --comb-c-penalty=3 --comb-k-penalty=1 --comb-s-penalty=5 --subvarsup=false --e-call-point=0.15")), (* sh8_b.comb.sh (modified) *) |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
476 |
((1, false, false, 1024, mashN), (format, "mono_native_higher_fool", keep_lamsN, false, "--mode=ho-pragmatic --boolean-reasoning=simpl-only --ho-unif-max-depth=0 --ho-prim-enum=none -q \"2|prefer-ho-steps|conjecture-relative-e(0.1,0.5,100,100,100,100,1.5,1.5,1.5)\" -q \"1|prefer-sos|pnrefined(1,1,1,2,2,2,0.5)\" -q \"2|prefer-ground|default\" -q \"2|prefer-empty-trail|conjecture-relative-e(0.1,0.5,100,100,100,100,1.5,1.5,1.5)\" -q \"1|prefer-processed|fifo\" --select=bb+e-selection7 --ho-pattern-decider=false --ho-fixpoint-decider=true --ho-solid-decider=false --sine=150 --sine-tolerance=2 --sine-depth-max=3 --sine-depth-min=1 --prec-gen-fun=invfreqhack --lazy-cnf=true --lazy-cnf-kind=simp --lazy-cnf-renaming-threshold=2 --fluid-log-hoist=false --tptp-def-as-rewrite --rewrite-before-cnf=true --ho-prim-enum=eq --ho-prim-enum-add-var=true --ho-prim-max=1 --ho-prim-enum-early-bird=true --avatar=eager --split-only-ground=true")), (* sh5_add_var_l_av.sh *) |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
477 |
((1, false, false, 256, mashN), (format, "mono_native_higher_fool", keep_lamsN, false, "--mode=ho-pragmatic --max-inferences=4 --ho-max-app-projections=1 --ho-max-elims=0 --ho-max-rigid-imitations=2 --ho-max-identifications=0 --ho-unif-max-depth=2 --max-inferences=3 --boolean-reasoning=bool-hoist --bool-select=LO --ext-rules=off --kbo-weight-fun=lambda-def-invfreqrank --ho-prim-enum=none --ho-unif-level=pragmatic-framework -q \"1|prefer-sos|conjecture-relative-var(1.01,s,f)\" -q \"4|const|conjecture-relative-var(1.05,l,f)\" -q \"1|prefer-processed|fifo\" -q \"1|prefer-non-goals|conjecture-relative-var(1.02,l,f)\" -q \"4|prefer-sos|pnrefined(3,2,3,2,2,1.5,2)\" --ho-elim-leibniz=1 --ho-fixpoint-decider=true --ho-pattern-decider=true --ho-solid-decider=true --select=e-selection2 --solve-formulas=true --lambdasup=0 --try-e=\"$E_HOME/eprover\" --tmp-dir=\"$ISABELLE_TMP_PREFIX\" --e-timeout=3 --e-max-derived=48 --e-encode-lambdas=lift --presaturate=true --prec-gen-fun=invfreq --e-call-point=0.2 --e-auto=true --sine-trim-implications=true")), (* sh10_e_lift.sh *) |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
478 |
((1, false, false, 512, mepoN), (format, "mono_native_higher_fool", keep_lamsN, false, "-nc --tptp-def-as-rewrite --mode=ho-pragmatic --max-inferences=4 --ho-max-app-projections=1 --ho-max-elims=0 --ho-max-rigid-imitations=1 --ho-max-identifications=0 --ho-unif-max-depth=2 --boolean-reasoning=simpl-only --ext-rules=off --kbo-weight-fun=lambda-def-const --ho-prim-enum=neg --ho-prim-enum-early-bird=true --tptp-def-as-rewrite --rewrite-before-cnf=true --ho-unif-level=pragmatic-framework --ho-unif-max-depth=1 --sine=50 --sine-tolerance=1.0 --sine-depth-max=3 --sine-depth-min=1 --sine-trim-implications=true -q \"4|prefer-sos|orient-lmax(2,1,2,1,1)\" -q \"4|defer-sos|conjecture-relative-var(1,s,f)\" -q \"3|const|default\" -q \"1|prefer-processed|fifo\" --ho-elim-leibniz=1 --ho-fixpoint-decider=true --ho-pattern-decider=true --ho-solid-decider=false --select=e-selection10 --solve-formulas=true --sup-at-vars=false --sup-at-var-headed=false --sup-from-var-headed=true --lazy-cnf=true --lazy-cnf-kind=simp --lazy-cnf-renaming-threshold=4 --try-e=\"$E_HOME/eprover\" --tmp-dir=\"$ISABELLE_TMP_PREFIX\" --e-timeout=2 --e-auto=true --e-max-derived=32 --e-encode-lambdas=lift --scan-clause-ac=false --presaturate=true --e-call-point=0.16")), (* sh5_shallow_sine.sh *) |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
479 |
((1, false, false, 512, mashN), (format, "mono_native_higher_fool", keep_lamsN, false, "-nc --tptp-def-as-rewrite --rewrite-before-cnf=true --mode=ho-competitive --boolean-reasoning=bool-hoist --ext-rules=off --recognize-injectivity=true --ho-unif-level=full-framework -q \"4|prefer-goals|pnrefined(1,1,1,2,2,2,0.5)\" -q \"1|prefer-fo|conjecture-relative-var(1.02,l,f)\" -q \"1|prefer-non-goals|conjecture-relative-var(1,l,f)\" -q \"4|prefer-ho-steps|conjecture-relative-var(1.01,s,f)\" -q \"1|prefer-processed|fifo\" --select=bb+ho-selection --scan-clause-ac=false --kbo-weight-fun=invfreqrank --fluidsup=true --boolean-reasoning=bool-hoist --fluid-log-hoist=true --fluid-hoist=true --ite-axioms=true --lazy-cnf=true --ho-solid-decider=true --ho-fixpoint-decider=true --bool-select=\"sel1(pos_ctx)\" --try-e=\"$E_HOME/eprover\" --tmp-dir=\"$ISABELLE_TMP_PREFIX\" --e-timeout=2 --e-auto=true --e-call-point=0.35 --avatar=off --e-max-derived=50")), (* sh5_e_short1.sh *) |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
480 |
((1, false, false, 1024, meshN), (format, "mono_native_higher_fool", keep_lamsN, false, "--mode=ho-pragmatic --tptp-def-as-rewrite --rewrite-before-cnf=true --boolean-reasoning=simpl-only --select=e-selection12 --prec-gen-fun=invfreq_conj --ord=lambda_kbo --ho-unif-level=full-framework --ho-pattern-decider=true --ho-solid-decider=false --ho-fixpoint-decider=true --lazy-cnf=true --lazy-cnf-kind=simp --lazy-cnf-renaming-threshold=1 --sine=100 --sine-depth-min=1 --sine-depth-max=5 --sine-tolerance=1.5 -q \"1|prefer-sos|default\" -q \"8|prefer-processed|conjecture-relative-var(1,l,f)\" -q \"1|prefer-processed|fifo\" -q \"1|prefer-lambdas|explore\" -q \"1|prefer-non-goals|explore\" -q \"1|prefer-processed|conjecture-relative-var(1,s,f)\" --kbo-weight-fun=arity0")), (* sh5_32.sh *) |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
481 |
((1, false, false, 256, meshN), (format, "mono_native_higher_fool", keep_lamsN, false, "--mode=ho-pragmatic --boolean-reasoning=simpl-only --ho-unif-max-depth=2 --ho-max-app-projections=0 --ho-max-rigid-imitations=1 --ho-max-elims=0 --ho-max-identifications=1 --max-inferences=3 --ext-rules=off --recognize-injectivity=true --ho-prim-enum=none --ho-choice-inst=true -q \"3|prefer-fo|conjecture-relative-var(1.02,l,f)\" -q \"3|prefer-sos|pnrefined(1,1,1,2,2,2,0.5)\" -q \"2|prefer-ground|orient-lmax(2,1,2,1,1)\" -q \"1|prefer-processed|fifo\" --select=MaxGoalNS --sine=60 --sine-tolerance=1.5 --sine-depth-max=3 --sine-depth-min=1 --prec-gen-fun=invfreqhack --lazy-cnf=true --lazy-cnf-kind=simp --lazy-cnf-renaming-threshold=3 --kbo-weight-fun-from-precedence=true --kbo-weight-fun-from-precedence-rank=5 --trigger-bool-inst=1 --avatar=lazy --tptp-def-as-rewrite --rewrite-before-cnf=true --sup-from-var-headed=false --sup-at-vars=false")), (* sh5_sh4.sh *) |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
482 |
((1, false, false, 512, mashN), (format, "mono_native_higher_fool", keep_lamsN, false, "--tptp-def-as-rewrite --rewrite-before-cnf=true --mode=lambda-free-intensional --check-lambda-free=false --boolean-reasoning=simpl-only --post-cnf-lambda-lifting=true --ext-rules=off --ho-prim-enum=none --recognize-injectivity=true --no-max-vars --select=e-selection8 --prec-gen-fun=invfreq --kbo-weight-fun=invfreqrank --kbo-const-weight=2 --ord=lambdafree_kbo --ignore-orphans=true -q \"1|prefer-sos|conjecture-relative-struct(1.5,3.5,2,3)\" -q \"1|const|conj_pref_weight(0.5,100,0.2,0.2,4)\" -q \"1|prefer-processed|fifo\" -q \"1|prefer-non-goals|conjecture-relative-e(0.3,0.25,100,100,100,100,1.5,1.5,1)\" -q \"1|prefer-sos|pnrefined(3,2,3,2,2,1.5,2)\" --lazy-cnf=true --lazy-cnf-renaming-threshold=2")), (* sh5_lifting2.sh *) |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
483 |
((1, false, false, 512, meshN), (format, "mono_native_higher_fool", keep_lamsN, false, "-nc --tptp-def-as-rewrite --rewrite-before-cnf=true --tptp-rewrite-formulas-only=true --mode=ho-pragmatic --boolean-reasoning=simpl-only --ext-rules=ext-family --ext-rules-max-depth=1 --ho-prim-enum=neg --ho-prim-max=1 --recognize-injectivity=true --ho-elim-leibniz=1 --ho-unif-level=pragmatic-framework --no-max-vars -q \"1|prefer-sos|conjecture-relative-var(1.02,l,f)\" -q \"4|const|conjecture-relative-var(1,s,f)\" -q \"1|prefer-processed|fifo\" -q \"1|prefer-non-goals|conjecture-relative-var(1,l,f)\" -q \"4|prefer-sos|pnrefined(2,1,1,1,2,2,2)\" --select=e-selection7 --ho-choice-inst=true --try-e=\"$E_HOME/eprover\" --tmp-dir=\"$ISABELLE_TMP_PREFIX\" --e-timeout=2 --e-auto=true --sine=50 --sine-tolerance=10 --sine-depth-max=5 --sine-depth-min=1 --e-max-derived=64 --e-encode-lambdas=lift --scan-clause-ac=false --prec-gen-fun=invfreq_conj --ord=lambda_kbo --solid-subsumption=true --ignore-orphans=true --e-call-point=0.2")), (* sh5_noforms.sh *) |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
484 |
((1, false, false, 1024, meshN), (format, "mono_native_higher_fool", keep_lamsN, false, "--mode=ho-pragmatic --tptp-def-as-rewrite --rewrite-before-cnf=true --max-inferences=4 --ho-unif-max-depth=3 --ho-max-elims=0 --ho-max-app-projections=1 --ho-max-identifications=0 --ho-max-rigid-imitations=2 --ho-unif-level=pragmatic-framework --boolean-reasoning=simpl-only --kbo-weight-fun=freqrank --ext-rules=ext-family --ext-rules-max-depth=2 --ho-prim-enum=eq --ho-prim-max=2 --interpret-bool-funs=false -q \"2|prefer-goals|default\" -q \"8|prefer-sos|conjecture-relative-var(1,l,f)\" -q \"1|prefer-processed|fifo\" -q \"1|prefer-lambdas|explore\" -q \"1|prefer-non-goals|explore\" -q \"1|prefer-processed|conjecture-relative-var(1,s,f)\" --recognize-injectivity=true --ho-selection-restriction=none --select=ho-selection2 --solve-formulas=true")), (* sh8_old_zip1.sh *) |
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
485 |
((1, false, false, 256, mashN), (format, "mono_native_higher_fool", keep_lamsN, false, "--mode=ho-pragmatic --boolean-reasoning=simpl-only --tptp-def-as-rewrite --rewrite-before-cnf=true --kbo-weight-fun=freqrank -q \"1|prefer-sos|default\" -q \"1|prefer-goals|conjecture-relative-var(1.03,s,f)\" -q \"1|prefer-non-goals|default\" -q \"5|const|conjecture-relative-var(1.01,l,f)\" -q \"1|prefer-processed|fifo\" -q \"1|const|conjecture-relative-var(1.05,l,f)\" -q \"1|defer-sos|conjecture-relative-var(1.1,s,f)\" --select=e-selection9 --recognize-injectivity=true --ho-choice-inst=false --ho-selection-restriction=none --sine=50 --sine-tolerance=3 --sine-depth-max=3 --sine-depth-min=1 --eq-encode=true --avatar=eager --sine-trim-implications=true"))], (* sh5_sh.eqenc.sh *) |
75026 | 486 |
good_max_mono_iters = default_max_mono_iters, |
487 |
good_max_new_mono_instances = default_max_new_mono_instances} |
|
73375 | 488 |
end |
57154 | 489 |
|
490 |
val zipperposition = (zipperpositionN, fn () => zipperposition_config) |
|
491 |
||
492 |
||
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
493 |
(* Remote ATP invocation via SystemOnTPTP *) |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
494 |
|
73426
bd8bce50b9d4
use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents:
73375
diff
changeset
|
495 |
val no_remote_systems = {url = "", systems = [] : string list} |
bd8bce50b9d4
use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents:
73375
diff
changeset
|
496 |
val remote_systems = Synchronized.var "atp_remote_systems" no_remote_systems |
31835 | 497 |
|
49984 | 498 |
fun get_remote_systems () = |
73426
bd8bce50b9d4
use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents:
73375
diff
changeset
|
499 |
Timeout.apply (seconds 10.0) SystemOnTPTP.list_systems () |
bd8bce50b9d4
use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents:
73375
diff
changeset
|
500 |
handle ERROR msg => (warning msg; no_remote_systems) |
bd8bce50b9d4
use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents:
73375
diff
changeset
|
501 |
| Timeout.TIMEOUT _ => no_remote_systems |
31835 | 502 |
|
49984 | 503 |
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
|
504 |
find_first (String.isPrefix (name ^ "---")) systems |
49984 | 505 |
| 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
|
506 |
case find_first (String.isPrefix (name ^ "---" ^ version)) systems of |
49984 | 507 |
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
|
508 |
| 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
|
509 |
|
49984 | 510 |
fun get_remote_system name versions = |
73426
bd8bce50b9d4
use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents:
73375
diff
changeset
|
511 |
Synchronized.change_result remote_systems (fn remote => |
bd8bce50b9d4
use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents:
73375
diff
changeset
|
512 |
(if #url remote <> SystemOnTPTP.get_url () orelse null (#systems remote) |
bd8bce50b9d4
use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents:
73375
diff
changeset
|
513 |
then get_remote_systems () else remote) |> ` #systems) |
bd8bce50b9d4
use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents:
73375
diff
changeset
|
514 |
|> `(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
|
515 |
|
49984 | 516 |
fun the_remote_system name versions = |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
517 |
(case get_remote_system name versions of |
42955 | 518 |
(SOME sys, _) => sys |
63692 | 519 |
| (NONE, []) => error "SystemOnTPTP is currently not available" |
42955 | 520 |
| (NONE, syss) => |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
521 |
(case syss |> filter_out (String.isPrefix "%") |> filter_out (curry (op =) "") of |
63692 | 522 |
[] => error "SystemOnTPTP is currently not available" |
523 |
| [msg] => error ("SystemOnTPTP is currently not available: " ^ msg) |
|
46480 | 524 |
| syss => |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
525 |
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
|
526 |
commas_quote syss ^ ".)"))) |
31835 | 527 |
|
72174 | 528 |
val max_remote_secs = 1000 (* give Geoff Sutcliffe's servers a break *) |
41148 | 529 |
|
73435
1cc848548f21
invoke remote ATP via SystemOnTPTP.run_systems from Isabelle/Scala (without perl);
wenzelm
parents:
73432
diff
changeset
|
530 |
val isabelle_scala_function = (["SCALA_HOME"], ["bin/scala"]) |
1cc848548f21
invoke remote ATP via SystemOnTPTP.run_systems from Isabelle/Scala (without perl);
wenzelm
parents:
73432
diff
changeset
|
531 |
|
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
532 |
fun remote_config system_name system_versions proof_delims known_failures prem_role |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
533 |
generate_isabelle_info good_slice = |
73435
1cc848548f21
invoke remote ATP via SystemOnTPTP.run_systems from Isabelle/Scala (without perl);
wenzelm
parents:
73432
diff
changeset
|
534 |
{exec = isabelle_scala_function, |
75341 | 535 |
arguments = fn _ => fn _ => fn command => fn timeout => fn problem => |
73435
1cc848548f21
invoke remote ATP via SystemOnTPTP.run_systems from Isabelle/Scala (without perl);
wenzelm
parents:
73432
diff
changeset
|
536 |
[the_remote_system system_name system_versions, |
1cc848548f21
invoke remote ATP via SystemOnTPTP.run_systems from Isabelle/Scala (without perl);
wenzelm
parents:
73432
diff
changeset
|
537 |
Isabelle_System.absolute_path problem, |
73568
bdba138d462d
clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents:
73437
diff
changeset
|
538 |
command, string_of_int (Int.min (max_remote_secs, to_secs 1 timeout) * 1000)], |
42962
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
blanchet
parents:
42955
diff
changeset
|
539 |
proof_delims = union (op =) tstp_proof_delims proof_delims, |
73436
e92f2e44e4d8
removed spurious references to perl / libwww-perl;
wenzelm
parents:
73435
diff
changeset
|
540 |
known_failures = known_failures @ known_says_failures, |
47976 | 541 |
prem_role = prem_role, |
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
542 |
generate_isabelle_info = generate_isabelle_info, |
75026 | 543 |
good_slices = fn ctxt => [good_slice ctxt], |
544 |
good_max_mono_iters = default_max_mono_iters, |
|
545 |
good_max_new_mono_instances = default_max_new_mono_instances} : atp_config |
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42332
diff
changeset
|
546 |
|
75026 | 547 |
fun remotify_config system_name system_versions good_slice |
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
548 |
({proof_delims, known_failures, prem_role, generate_isabelle_info, ...} : atp_config) = |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
549 |
remote_config system_name system_versions proof_delims known_failures prem_role |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
550 |
generate_isabelle_info good_slice |
38023 | 551 |
|
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
552 |
fun remote_atp name system_name system_versions proof_delims known_failures prem_role |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
553 |
generate_isabelle_info good_slice = |
58084 | 554 |
(remote_prefix ^ name, fn () => |
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
555 |
remote_config system_name system_versions proof_delims known_failures prem_role |
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
556 |
generate_isabelle_info good_slice) |
75026 | 557 |
fun remotify_atp (name, config) system_name system_versions good_slice = |
558 |
(remote_prefix ^ name, remotify_config system_name system_versions good_slice o config) |
|
28592 | 559 |
|
57269
1df6f747f164
changed type encoding for new Waldmeister, to trigger filtering of 'dangerous' lemmas
blanchet
parents:
57265
diff
changeset
|
560 |
fun gen_remote_waldmeister name type_enc = |
57265 | 561 |
remote_atp name "Waldmeister" ["710"] tstp_proof_delims |
562 |
([(OutOfResources, "Too many function symbols"), |
|
563 |
(Inappropriate, "**** Unexpected end of file."), |
|
564 |
(Crashed, "Unrecoverable Segmentation Fault")] |
|
565 |
@ known_szs_status_failures) |
|
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
566 |
Hypothesis false |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
76301
diff
changeset
|
567 |
(K ((1000 (* infinity *), false, false, 50, meshN), (CNF_UEQ, type_enc, combsN, false, "")) (* FUDGE *)) |
57264 | 568 |
|
52094 | 569 |
val remote_agsyhol = |
570 |
remotify_atp agsyhol "agsyHOL" ["1.0", "1"] |
|
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
76301
diff
changeset
|
571 |
(K ((1000 (* infinity *), false, false, 60, meshN), (THF (Monomorphic, {with_ite = false, with_let = false}, THF_Without_Choice), "mono_native_higher", keep_lamsN, false, "")) (* FUDGE *)) |
70937 | 572 |
val remote_alt_ergo = |
573 |
remotify_atp alt_ergo "Alt-Ergo" ["0.95.2"] |
|
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
76301
diff
changeset
|
574 |
(K ((1000 (* infinity *), false, false, 250, meshN), (TF1, "poly_native", keep_lamsN, false, "")) (* FUDGE *)) |
43500
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents:
43497
diff
changeset
|
575 |
val remote_e = |
63768 | 576 |
remotify_atp e "E" ["2.0", "1.9.1", "1.8"] |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
76301
diff
changeset
|
577 |
(K ((1000 (* infinity *), false, false, 750, meshN), (TF0, "mono_native", combsN, false, "")) (* FUDGE *)) |
48700 | 578 |
val remote_iprover = |
52094 | 579 |
remotify_atp iprover "iProver" ["0.99"] |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
76301
diff
changeset
|
580 |
(K ((1000 (* infinity *), false, false, 150, meshN), (FOF, "mono_guards??", liftingN, false, "")) (* FUDGE *)) |
44099 | 581 |
val remote_leo2 = |
52094 | 582 |
remotify_atp leo2 "LEO-II" ["1.5.0", "1.4", "1.3", "1.2", "1"] |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
76301
diff
changeset
|
583 |
(K ((1000 (* infinity *), false, false, 40, meshN), (THF (Monomorphic, {with_ite = false, with_let = false}, THF_Without_Choice), "mono_native_higher", liftingN, false, "")) (* FUDGE *)) |
67021
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66544
diff
changeset
|
584 |
val remote_leo3 = |
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66544
diff
changeset
|
585 |
remotify_atp leo3 "Leo-III" ["1.1"] |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
76301
diff
changeset
|
586 |
(K ((1000 (* infinity *), false, false, 150, meshN), (THF (Polymorphic, {with_ite = false, with_let = false}, THF_Without_Choice), "poly_native_higher", keep_lamsN, false, "")) (* FUDGE *)) |
57269
1df6f747f164
changed type encoding for new Waldmeister, to trigger filtering of 'dangerous' lemmas
blanchet
parents:
57265
diff
changeset
|
587 |
val remote_waldmeister = gen_remote_waldmeister waldmeisterN "raw_mono_tags??" |
70940
ce3a05ad07b7
added support for Zipperposition on SystemOnTPTP
blanchet
parents:
70939
diff
changeset
|
588 |
val remote_zipperposition = |
74468 | 589 |
remotify_atp zipperposition "Zipperpin" ["2.1", "2.0"] |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
76301
diff
changeset
|
590 |
(K ((1000 (* infinity *), false, false, 512, meshN), (THF (Monomorphic, {with_ite = false, with_let = false}, THF_Without_Choice), "mono_native_higher", keep_lamsN, false, "")) (* FUDGE *)) |
72588 | 591 |
|
592 |
||
593 |
(* Dummy prover *) |
|
594 |
||
595 |
fun dummy_config prem_role format type_enc uncurried_aliases : atp_config = |
|
73374 | 596 |
{exec = (["ISABELLE_ATP"], ["scripts/dummy_atp"]), |
75341 | 597 |
arguments = K (K (K (K (K [])))), |
72588 | 598 |
proof_delims = [], |
599 |
known_failures = known_szs_status_failures, |
|
600 |
prem_role = prem_role, |
|
76301
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
blanchet
parents:
76183
diff
changeset
|
601 |
generate_isabelle_info = false, |
75026 | 602 |
good_slices = |
82231
cbe937aa5e90
merged, resolving conflicts in src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML due to clones bb2ea9e80c33 + 62c039ce397c;
wenzelm
parents:
82204
diff
changeset
|
603 |
K [((2, false, false, 256, "mepo"), (format, type_enc, |
75026 | 604 |
if is_format_higher_order format then keep_lamsN else combsN, uncurried_aliases, ""))], |
605 |
good_max_mono_iters = default_max_mono_iters, |
|
606 |
good_max_new_mono_instances = default_max_new_mono_instances} |
|
72588 | 607 |
|
74894
a64a8f7d593e
refactored $ite and $let configuration and added dummy_thf_reduced prover
desharna
parents:
74890
diff
changeset
|
608 |
val dummy_fof = |
75024
114eb0af123d
simplified 'best_slice' data structure and made minor changes to slices
blanchet
parents:
74999
diff
changeset
|
609 |
(dummy_fofN, fn () => dummy_config Hypothesis FOF "mono_guards??" false) |
74117 | 610 |
|
74894
a64a8f7d593e
refactored $ite and $let configuration and added dummy_thf_reduced prover
desharna
parents:
74890
diff
changeset
|
611 |
val dummy_tfx = |
75024
114eb0af123d
simplified 'best_slice' data structure and made minor changes to slices
blanchet
parents:
74999
diff
changeset
|
612 |
(dummy_tfxN, fn () => dummy_config Hypothesis TX1 "poly_native_fool" false) |
74117 | 613 |
|
74894
a64a8f7d593e
refactored $ite and $let configuration and added dummy_thf_reduced prover
desharna
parents:
74890
diff
changeset
|
614 |
val dummy_thf = |
75024
114eb0af123d
simplified 'best_slice' data structure and made minor changes to slices
blanchet
parents:
74999
diff
changeset
|
615 |
(dummy_thfN, fn () => dummy_config Hypothesis TH1 "poly_native_higher" false) |
38454
9043eefe8d71
detect old Vampire and give a nicer error message
blanchet
parents:
38433
diff
changeset
|
616 |
|
74894
a64a8f7d593e
refactored $ite and $let configuration and added dummy_thf_reduced prover
desharna
parents:
74890
diff
changeset
|
617 |
val dummy_thf_reduced = |
a64a8f7d593e
refactored $ite and $let configuration and added dummy_thf_reduced prover
desharna
parents:
74890
diff
changeset
|
618 |
let |
74895
b605ebd87a47
proper polymorphism for TH1 format in Sledgehammer
desharna
parents:
74894
diff
changeset
|
619 |
val format = THF (Polymorphic, {with_ite = false, with_let = false}, THF_Without_Choice) |
74894
a64a8f7d593e
refactored $ite and $let configuration and added dummy_thf_reduced prover
desharna
parents:
74890
diff
changeset
|
620 |
val config = dummy_config Hypothesis format "poly_native_higher" false |
a64a8f7d593e
refactored $ite and $let configuration and added dummy_thf_reduced prover
desharna
parents:
74890
diff
changeset
|
621 |
in (dummy_thfN ^ "_reduced", fn () => config) end |
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
622 |
|
38454
9043eefe8d71
detect old Vampire and give a nicer error message
blanchet
parents:
38433
diff
changeset
|
623 |
(* Setup *) |
9043eefe8d71
detect old Vampire and give a nicer error message
blanchet
parents:
38433
diff
changeset
|
624 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
625 |
fun add_atp (name, config) thy = |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
626 |
Data.map (Symtab.update_new (name, (config, stamp ()))) thy |
63692 | 627 |
handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name) |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
628 |
|
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
629 |
fun get_atp thy name = |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
630 |
fst (the (Symtab.lookup (Data.get thy) name)) |
63692 | 631 |
handle Option.Option => error ("Unknown ATP: " ^ name) |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
632 |
|
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
633 |
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
|
634 |
let val {exec, ...} = get_atp thy name () in |
73374 | 635 |
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
|
636 |
end |
36371
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
637 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
638 |
fun refresh_systems_on_tptp () = |
49984 | 639 |
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
|
640 |
|
75036 | 641 |
val local_atps = |
75038 | 642 |
[agsyhol, alt_ergo, e, iprover, leo2, leo3, satallax, spass, vampire, zipperposition] |
75036 | 643 |
val remote_atps = |
644 |
[remote_agsyhol, remote_alt_ergo, remote_e, remote_iprover, remote_leo2, remote_leo3, |
|
75465
d9b23902692d
excluded dummy ATPs from Sledgehammer's default provers
desharna
parents:
75433
diff
changeset
|
645 |
remote_waldmeister, remote_zipperposition] |
d9b23902692d
excluded dummy ATPs from Sledgehammer's default provers
desharna
parents:
75433
diff
changeset
|
646 |
val dummy_atps = |
d9b23902692d
excluded dummy ATPs from Sledgehammer's default provers
desharna
parents:
75433
diff
changeset
|
647 |
[dummy_fof, dummy_tfx, dummy_thf, dummy_thf_reduced] |
d9b23902692d
excluded dummy ATPs from Sledgehammer's default provers
desharna
parents:
75433
diff
changeset
|
648 |
val non_dummy_atps = local_atps @ remote_atps |
d9b23902692d
excluded dummy ATPs from Sledgehammer's default provers
desharna
parents:
75433
diff
changeset
|
649 |
val all_atps = non_dummy_atps @ dummy_atps |
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
|
650 |
|
75465
d9b23902692d
excluded dummy ATPs from Sledgehammer's default provers
desharna
parents:
75433
diff
changeset
|
651 |
val _ = Theory.setup (fold add_atp all_atps) |
35867 | 652 |
|
75036 | 653 |
val local_atps = map fst local_atps |
654 |
val remote_atps = map fst remote_atps |
|
75465
d9b23902692d
excluded dummy ATPs from Sledgehammer's default provers
desharna
parents:
75433
diff
changeset
|
655 |
val dummy_atps = map fst dummy_atps |
d9b23902692d
excluded dummy ATPs from Sledgehammer's default provers
desharna
parents:
75433
diff
changeset
|
656 |
val non_dummy_atps = map fst non_dummy_atps |
d9b23902692d
excluded dummy ATPs from Sledgehammer's default provers
desharna
parents:
75433
diff
changeset
|
657 |
val all_atps = map fst all_atps |
75036 | 658 |
|
28592 | 659 |
end; |