| author | desharna | 
| Thu, 08 Oct 2020 16:36:00 +0200 | |
| changeset 72399 | f8900a5ad4a7 | 
| parent 72174 | 585b877df698 | 
| permissions | -rw-r--r-- | 
| 38047 | 1 | (* Title: HOL/Tools/ATP/atp_systems.ML | 
| 28592 | 2 | Author: Fabian Immler, TU Muenchen | 
| 36371 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 3 | Author: Jasmin Blanchette, TU Muenchen | 
| 28592 | 4 | |
| 36376 | 5 | Setup for supported ATPs. | 
| 28592 | 6 | *) | 
| 7 | ||
| 36376 | 8 | signature ATP_SYSTEMS = | 
| 28592 | 9 | sig | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 10 | type term_order = ATP_Problem.term_order | 
| 45301 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 11 | type atp_format = ATP_Problem.atp_format | 
| 53586 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 blanchet parents: 
53515diff
changeset | 12 | type atp_formula_role = ATP_Problem.atp_formula_role | 
| 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 blanchet parents: 
53515diff
changeset | 13 | type atp_failure = ATP_Proof.atp_failure | 
| 38023 | 14 | |
| 51011 | 15 | type slice_spec = (int * string) * atp_format * string * string * bool | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 16 | type atp_config = | 
| 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: 
57636diff
changeset | 17 |     {exec : bool -> string list * string list,
 | 
| 
dc5e1b1db9ba
avoid 'eproof' and 'eproof_ram' scripts if possible (i.e. if 'eprover' can produce reasonable enough proofs for one-liner reconstruction)
 blanchet parents: 
57636diff
changeset | 18 | arguments : Proof.context -> bool -> string -> Time.time -> string -> | 
| 
dc5e1b1db9ba
avoid 'eproof' and 'eproof_ram' scripts if possible (i.e. if 'eprover' can produce reasonable enough proofs for one-liner reconstruction)
 blanchet parents: 
57636diff
changeset | 19 | term_order * (unit -> (string * int) list) * (unit -> (string * real) list) -> string, | 
| 42578 
1eaf4d437d4c
define type system in ATP module so that ATPs can suggest a type system
 blanchet parents: 
42577diff
changeset | 20 | proof_delims : (string * string) list, | 
| 53586 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 blanchet parents: 
53515diff
changeset | 21 | known_failures : (atp_failure * string) list, | 
| 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 blanchet parents: 
53515diff
changeset | 22 | prem_role : atp_formula_role, | 
| 48716 
1d2a12bb0640
stop distinguishing between complete and incomplete slices, since this is very fragile and has hardly any useful semantics to users
 blanchet parents: 
48715diff
changeset | 23 | best_slices : Proof.context -> (real * (slice_spec * string)) list, | 
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 24 | best_max_mono_iters : int, | 
| 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 25 | best_max_new_mono_instances : int} | 
| 38023 | 26 | |
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 27 | val default_max_mono_iters : int | 
| 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 28 | val default_max_new_mono_instances : int | 
| 44099 | 29 | val force_sos : bool Config.T | 
| 47032 | 30 | val term_order : string Config.T | 
| 43566 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 blanchet parents: 
43529diff
changeset | 31 | val e_smartN : string | 
| 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 blanchet parents: 
43529diff
changeset | 32 | val e_autoN : string | 
| 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 blanchet parents: 
43529diff
changeset | 33 | val e_fun_weightN : string | 
| 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 blanchet parents: 
43529diff
changeset | 34 | val e_sym_offset_weightN : string | 
| 47032 | 35 | val e_selection_heuristic : string Config.T | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 36 | val e_default_fun_weight : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 37 | val e_fun_weight_base : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 38 | val e_fun_weight_span : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 39 | val e_default_sym_offs_weight : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 40 | val e_sym_offs_weight_base : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 41 | val e_sym_offs_weight_span : real Config.T | 
| 50950 | 42 | val spass_H1SOS : string | 
| 43 | val spass_H2 : string | |
| 44 | val spass_H2LR0LT0 : string | |
| 45 | val spass_H2NuVS0 : string | |
| 46 | val spass_H2NuVS0Red2 : string | |
| 47 | val spass_H2SOS : string | |
| 68563 | 48 | val is_vampire_noncommercial_license_accepted : unit -> bool option | 
| 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: 
57636diff
changeset | 49 | val remote_atp : string -> string -> string list -> (string * string) list -> | 
| 
dc5e1b1db9ba
avoid 'eproof' and 'eproof_ram' scripts if possible (i.e. if 'eprover' can produce reasonable enough proofs for one-liner reconstruction)
 blanchet parents: 
57636diff
changeset | 50 | (atp_failure * string) list -> atp_formula_role -> (Proof.context -> slice_spec * string) -> | 
| 
dc5e1b1db9ba
avoid 'eproof' and 'eproof_ram' scripts if possible (i.e. if 'eprover' can produce reasonable enough proofs for one-liner reconstruction)
 blanchet parents: 
57636diff
changeset | 51 | string * (unit -> atp_config) | 
| 47606 
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
 blanchet parents: 
47506diff
changeset | 52 | val add_atp : string * (unit -> atp_config) -> theory -> theory | 
| 
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
 blanchet parents: 
47506diff
changeset | 53 | val get_atp : theory -> string -> (unit -> atp_config) | 
| 41727 
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
 blanchet parents: 
41725diff
changeset | 54 | val supported_atps : theory -> string list | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 55 | val is_atp_installed : theory -> string -> bool | 
| 35867 | 56 | val refresh_systems_on_tptp : unit -> unit | 
| 47055 
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
 blanchet parents: 
47053diff
changeset | 57 | val effective_term_order : Proof.context -> string -> term_order | 
| 28592 | 58 | end; | 
| 59 | ||
| 36376 | 60 | structure ATP_Systems : ATP_SYSTEMS = | 
| 28592 | 61 | struct | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 62 | |
| 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: 
42571diff
changeset | 63 | open ATP_Problem | 
| 39491 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39375diff
changeset | 64 | open ATP_Proof | 
| 46320 | 65 | open ATP_Problem_Generate | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 66 | |
| 52073 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 67 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 68 | (* ATP configuration *) | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 69 | |
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 70 | 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: 
53225diff
changeset | 71 | val default_max_new_mono_instances = 100 (* FUDGE *) | 
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 72 | |
| 51011 | 73 | type slice_spec = (int * string) * atp_format * string * string * bool | 
| 46409 
d4754183ccce
made option available to users (mostly for experiments)
 blanchet parents: 
46407diff
changeset | 74 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 75 | type atp_config = | 
| 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: 
57636diff
changeset | 76 |   {exec : bool -> string list * string list,
 | 
| 
dc5e1b1db9ba
avoid 'eproof' and 'eproof_ram' scripts if possible (i.e. if 'eprover' can produce reasonable enough proofs for one-liner reconstruction)
 blanchet parents: 
57636diff
changeset | 77 | arguments : Proof.context -> bool -> string -> Time.time -> string -> | 
| 
dc5e1b1db9ba
avoid 'eproof' and 'eproof_ram' scripts if possible (i.e. if 'eprover' can produce reasonable enough proofs for one-liner reconstruction)
 blanchet parents: 
57636diff
changeset | 78 | term_order * (unit -> (string * int) list) * (unit -> (string * real) list) -> string, | 
| 42578 
1eaf4d437d4c
define type system in ATP module so that ATPs can suggest a type system
 blanchet parents: 
42577diff
changeset | 79 | proof_delims : (string * string) list, | 
| 53586 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 blanchet parents: 
53515diff
changeset | 80 | known_failures : (atp_failure * string) list, | 
| 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 blanchet parents: 
53515diff
changeset | 81 | prem_role : atp_formula_role, | 
| 48716 
1d2a12bb0640
stop distinguishing between complete and incomplete slices, since this is very fragile and has hardly any useful semantics to users
 blanchet parents: 
48715diff
changeset | 82 | best_slices : Proof.context -> (real * (slice_spec * string)) list, | 
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 83 | best_max_mono_iters : int, | 
| 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 84 | best_max_new_mono_instances : int} | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 85 | |
| 42723 | 86 | (* "best_slices" must be found empirically, taking a wholistic approach since | 
| 51011 | 87 | the ATPs are run in parallel. Each slice has the format | 
| 88 | ||
| 89 | (time_frac, ((max_facts, fact_filter), format, type_enc, | |
| 90 | lam_trans, uncurried_aliases), extra) | |
| 91 | ||
| 92 | where | |
| 93 | ||
| 94 | time_frac = faction of the time available given to the slice (which should | |
| 95 | add up to 1.0) | |
| 96 | ||
| 97 | extra = extra information to the prover (e.g., SOS or no SOS). | |
| 42723 | 98 | |
| 99 | The last slice should be the most "normal" one, because it will get all the | |
| 43569 
b342cd125533
removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
 blanchet parents: 
43567diff
changeset | 100 | time available if the other slices fail early and also because it is used if | 
| 
b342cd125533
removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
 blanchet parents: 
43567diff
changeset | 101 | slicing is disabled (e.g., by the minimizer). *) | 
| 42710 
84fcce345b5d
further improved type system setup based on Judgment Days
 blanchet parents: 
42709diff
changeset | 102 | |
| 51011 | 103 | val mepoN = "mepo" | 
| 104 | val mashN = "mash" | |
| 105 | val meshN = "mesh" | |
| 106 | ||
| 52073 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 107 | val tstp_proof_delims = | 
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 108 |   [("% SZS output start CNFRefutation", "% SZS output end CNFRefutation"),
 | 
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 109 |    ("% SZS output start Refutation", "% SZS output end Refutation"),
 | 
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 110 |    ("% SZS output start Proof", "% SZS output end Proof")]
 | 
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 111 | |
| 38061 
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
 blanchet parents: 
38049diff
changeset | 112 | val known_perl_failures = | 
| 38094 | 113 | [(CantConnect, "HTTP error"), | 
| 114 | (NoPerl, "env: perl"), | |
| 38065 | 115 | (NoLibwwwPerl, "Can't locate HTTP")] | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 116 | |
| 45203 | 117 | fun known_szs_failures wrap = | 
| 118 | [(Unprovable, wrap "CounterSatisfiable"), | |
| 119 | (Unprovable, wrap "Satisfiable"), | |
| 120 | (GaveUp, wrap "GaveUp"), | |
| 121 | (GaveUp, wrap "Unknown"), | |
| 122 | (GaveUp, wrap "Incomplete"), | |
| 123 | (ProofMissing, wrap "Theorem"), | |
| 124 | (ProofMissing, wrap "Unsatisfiable"), | |
| 125 | (TimedOut, wrap "Timeout"), | |
| 126 | (Inappropriate, wrap "Inappropriate"), | |
| 127 | (OutOfResources, wrap "ResourceOut"), | |
| 128 | (OutOfResources, wrap "MemoryOut"), | |
| 129 | (Interrupted, wrap "Forced"), | |
| 130 | (Interrupted, wrap "User")] | |
| 131 | ||
| 132 | val known_szs_status_failures = known_szs_failures (prefix "SZS status ") | |
| 133 | val known_says_failures = known_szs_failures (prefix " says ") | |
| 134 | ||
| 38023 | 135 | structure Data = Theory_Data | 
| 136 | ( | |
| 47606 
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
 blanchet parents: 
47506diff
changeset | 137 | type T = ((unit -> atp_config) * stamp) Symtab.table | 
| 38023 | 138 | val empty = Symtab.empty | 
| 139 | val extend = I | |
| 46407 
30e9720cc0b9
optimization: slice caching in case two consecutive slices are nearly identical
 blanchet parents: 
46402diff
changeset | 140 | fun merge data : T = | 
| 
30e9720cc0b9
optimization: slice caching in case two consecutive slices are nearly identical
 blanchet parents: 
46402diff
changeset | 141 | Symtab.merge (eq_snd (op =)) data | 
| 63692 | 142 |     handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name)
 | 
| 38023 | 143 | ) | 
| 38017 
3ad3e3ca2451
move Sledgehammer-specific code out of "Sledgehammer_TPTP_Format"
 blanchet parents: 
38015diff
changeset | 144 | |
| 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: 
43850diff
changeset | 145 | 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: 
36064diff
changeset | 146 | |
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 147 | val sosN = "sos" | 
| 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 148 | val no_sosN = "no_sos" | 
| 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 149 | |
| 69593 | 150 | val force_sos = Attrib.setup_config_bool \<^binding>\<open>atp_force_sos\<close> (K false) | 
| 44099 | 151 | |
| 47032 | 152 | val smartN = "smart" | 
| 47073 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 blanchet parents: 
47055diff
changeset | 153 | (* val kboN = "kbo" *) | 
| 47032 | 154 | val lpoN = "lpo" | 
| 47034 
77da780ddd6b
implement term order attribute (for experiments)
 blanchet parents: 
47033diff
changeset | 155 | val xweightsN = "_weights" | 
| 
77da780ddd6b
implement term order attribute (for experiments)
 blanchet parents: 
47033diff
changeset | 156 | val xprecN = "_prec" | 
| 
77da780ddd6b
implement term order attribute (for experiments)
 blanchet parents: 
47033diff
changeset | 157 | val xsimpN = "_simp" (* SPASS-specific *) | 
| 47032 | 158 | |
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 159 | (* Possible values for "atp_term_order": | 
| 47049 | 160 | "smart", "(kbo|lpo)(_weights)?(_prec|_simp)?" *) | 
| 47032 | 161 | val term_order = | 
| 69593 | 162 | Attrib.setup_config_string \<^binding>\<open>atp_term_order\<close> (K smartN) | 
| 47032 | 163 | |
| 52073 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 164 | |
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 165 | (* agsyHOL *) | 
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 166 | |
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 167 | val agsyhol_config : atp_config = | 
| 52754 
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
 blanchet parents: 
52151diff
changeset | 168 |   {exec = K (["AGSYHOL_HOME"], ["agsyHOL"]),
 | 
| 52073 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 169 | arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => | 
| 68563 | 170 | "--proof --time-out " ^ string_of_int (to_secs 1 timeout) ^ " " ^ file_name, | 
| 52073 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 171 | proof_delims = tstp_proof_delims, | 
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 172 | known_failures = known_szs_status_failures, | 
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 173 | prem_role = Hypothesis, | 
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 174 | best_slices = | 
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 175 | (* FUDGE *) | 
| 70932 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 176 | K [(1.0, (((60, ""), THF (Monomorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN, false), ""))], | 
| 52073 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 177 | best_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *), | 
| 53515 
f5b678b155f6
adjusted number of generated monomorphic instances for new monomorphizer based on new evaluation (E, SPASS, Vampire)
 blanchet parents: 
53225diff
changeset | 178 | best_max_new_mono_instances = default_max_new_mono_instances} | 
| 52073 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 179 | |
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 180 | val agsyhol = (agsyholN, fn () => agsyhol_config) | 
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 181 | |
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 182 | |
| 46643 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 183 | (* Alt-Ergo *) | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 184 | |
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 185 | val alt_ergo_config : atp_config = | 
| 52754 
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
 blanchet parents: 
52151diff
changeset | 186 |   {exec = K (["WHY3_HOME"], ["why3"]),
 | 
| 50927 | 187 | arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => | 
| 68563 | 188 | "--format tptp --prover 'Alt-Ergo,0.95.2,' --timelimit " ^ string_of_int (to_secs 1 timeout) ^ | 
| 189 | " " ^ file_name, | |
| 46643 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 190 | proof_delims = [], | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 191 | known_failures = | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 192 | [(ProofMissing, ": Valid"), | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 193 | (TimedOut, ": Timeout"), | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 194 | (GaveUp, ": Unknown")], | 
| 47976 | 195 | prem_role = Hypothesis, | 
| 46643 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 196 | best_slices = fn _ => | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 197 | (* FUDGE *) | 
| 57547 
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
 blanchet parents: 
57293diff
changeset | 198 | [(1.0, (((100, ""), TFF Polymorphic, "poly_native", liftingN, false), ""))], | 
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 199 | best_max_mono_iters = default_max_mono_iters, | 
| 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 200 | best_max_new_mono_instances = default_max_new_mono_instances} | 
| 46643 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 201 | |
| 47646 | 202 | 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: 
46481diff
changeset | 203 | |
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 204 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 205 | (* E *) | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 206 | |
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 207 | val e_smartN = "smart" | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 208 | val e_autoN = "auto" | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 209 | val e_fun_weightN = "fun_weight" | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 210 | val e_sym_offset_weightN = "sym_offset_weight" | 
| 41725 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 211 | |
| 47032 | 212 | val e_selection_heuristic = | 
| 69593 | 213 | Attrib.setup_config_string \<^binding>\<open>atp_e_selection_heuristic\<close> (K e_smartN) | 
| 41770 | 214 | (* FUDGE *) | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 215 | val e_default_fun_weight = | 
| 69593 | 216 | Attrib.setup_config_real \<^binding>\<open>atp_e_default_fun_weight\<close> (K 20.0) | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 217 | val e_fun_weight_base = | 
| 69593 | 218 | Attrib.setup_config_real \<^binding>\<open>atp_e_fun_weight_base\<close> (K 0.0) | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 219 | val e_fun_weight_span = | 
| 69593 | 220 | Attrib.setup_config_real \<^binding>\<open>atp_e_fun_weight_span\<close> (K 40.0) | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 221 | val e_default_sym_offs_weight = | 
| 69593 | 222 | Attrib.setup_config_real \<^binding>\<open>atp_e_default_sym_offs_weight\<close> (K 1.0) | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 223 | val e_sym_offs_weight_base = | 
| 69593 | 224 | Attrib.setup_config_real \<^binding>\<open>atp_e_sym_offs_weight_base\<close> (K ~20.0) | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 225 | val e_sym_offs_weight_span = | 
| 69593 | 226 | Attrib.setup_config_real \<^binding>\<open>atp_e_sym_offs_weight_span\<close> (K 60.0) | 
| 41725 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 227 | |
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 228 | fun e_selection_heuristic_case heuristic fw sow = | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 229 | if heuristic = e_fun_weightN then fw | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 230 | else if heuristic = e_sym_offset_weightN then sow | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 231 |   else raise Fail ("unexpected " ^ quote heuristic)
 | 
| 41725 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 232 | |
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 233 | fun scaled_e_selection_weight ctxt heuristic w = | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 234 | w * Config.get ctxt (e_selection_heuristic_case heuristic | 
| 47029 | 235 | e_fun_weight_span e_sym_offs_weight_span) | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 236 | + Config.get ctxt (e_selection_heuristic_case heuristic | 
| 47029 | 237 | e_fun_weight_base e_sym_offs_weight_base) | 
| 41725 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 238 | |> Real.ceil |> signed_string_of_int | 
| 41313 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 blanchet parents: 
41269diff
changeset | 239 | |
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 240 | fun e_selection_weight_arguments ctxt heuristic sel_weights = | 
| 51631 | 241 | if heuristic = e_fun_weightN orelse heuristic = e_sym_offset_weightN then | 
| 43622 | 242 | (* supplied by Stephan Schulz *) | 
| 41314 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 243 | "--split-clauses=4 --split-reuse-defs --simul-paramod --forward-context-sr \ | 
| 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 244 | \--destructive-er-aggressive --destructive-er --presat-simplify \ | 
| 47505 
e33d957ae2bf
avoid option introduced in E 1.2 when invoking older versions of E
 blanchet parents: 
47499diff
changeset | 245 | \--prefer-initial-clauses -winvfreqrank -c1 -Ginvfreqconjmax -F1 \ | 
| 
e33d957ae2bf
avoid option introduced in E 1.2 when invoking older versions of E
 blanchet parents: 
47499diff
changeset | 246 | \--delete-bad-limit=150000000 -WSelectMaxLComplexAvoidPosPred -H'(4*" ^ | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 247 | e_selection_heuristic_case heuristic "FunWeight" "SymOffsetWeight" ^ | 
| 48376 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48232diff
changeset | 248 | "(SimulateSOS," ^ | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 249 | (e_selection_heuristic_case heuristic | 
| 47029 | 250 | e_default_fun_weight e_default_sym_offs_weight | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 251 | |> Config.get ctxt |> Real.ceil |> signed_string_of_int) ^ | 
| 41314 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 252 | ",20,1.5,1.5,1" ^ | 
| 47030 | 253 | (sel_weights () | 
| 47029 | 254 | |> map (fn (s, w) => "," ^ s ^ ":" ^ | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 255 | scaled_e_selection_weight ctxt heuristic w) | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 256 | |> implode) ^ | 
| 41314 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 257 | "),3*ConjectureGeneralSymbolWeight(PreferNonGoals,200,100,200,50,50,1,100,\ | 
| 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 258 | \1.5,1.5,1),1*Clauseweight(PreferProcessed,1,1,1),1*\ | 
| 57672 | 259 | \FIFOWeight(PreferProcessed))' " | 
| 51631 | 260 | else | 
| 57672 | 261 | "-xAuto " | 
| 41313 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 blanchet parents: 
41269diff
changeset | 262 | |
| 70939 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 263 | val e_ord_weights = map (fn (s, w) => s ^ ":" ^ string_of_int w) #> space_implode "," | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 264 | fun e_ord_precedence [_] = "" | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 265 | | e_ord_precedence info = info |> map fst |> space_implode "<" | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 266 | |
| 47039 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47038diff
changeset | 267 | fun e_term_order_info_arguments false false _ = "" | 
| 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47038diff
changeset | 268 | | e_term_order_info_arguments gen_weights gen_prec ord_info = | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 269 | let val ord_info = ord_info () in | 
| 57672 | 270 | (if gen_weights then "--order-weights='" ^ e_ord_weights ord_info ^ "' " else "") ^ | 
| 271 | (if gen_prec then "--precedence='" ^ e_ord_precedence ord_info ^ "' " else "") | |
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 272 | end | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 273 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 274 | val e_config : atp_config = | 
| 66363 | 275 |   {exec = fn _ => (["E_HOME"], ["eprover"]),
 | 
| 276 | arguments = fn ctxt => fn _ => fn heuristic => fn timeout => fn file_name => | |
| 57008 | 277 |      fn ({is_lpo, gen_weights, gen_prec, ...}, ord_info, sel_weights) =>
 | 
| 66363 | 278 | "--auto-schedule --tstp-in --tstp-out --silent " ^ | 
| 57672 | 279 | e_selection_weight_arguments ctxt heuristic sel_weights ^ | 
| 280 | e_term_order_info_arguments gen_weights gen_prec ord_info ^ | |
| 57636 
3ab503b04bdb
stick to external proofs when invoking E, because they are more detailed and do not merge steps
 blanchet parents: 
57547diff
changeset | 281 | "--term-ordering=" ^ (if is_lpo then "LPO4" else "KBO6") ^ " " ^ | 
| 48651 | 282 | "--cpu-limit=" ^ string_of_int (to_secs 2 timeout) ^ | 
| 66363 | 283 | " --proof-object=1 " ^ | 
| 284 | file_name, | |
| 52073 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 285 | proof_delims = | 
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 286 |      [("# SZS output start CNFRefutation", "# SZS output end CNFRefutation")] @
 | 
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 287 | tstp_proof_delims, | 
| 36265 
41c9e755e552
distinguish between the different ATP errors in the user interface;
 blanchet parents: 
36264diff
changeset | 288 | known_failures = | 
| 45203 | 289 | [(TimedOut, "Failure: Resource limit exceeded (time)"), | 
| 47972 | 290 | (TimedOut, "time limit exceeded")] @ | 
| 291 | known_szs_status_failures, | |
| 47976 | 292 | prem_role = Conjecture, | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 293 | best_slices = fn ctxt => | 
| 70939 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 294 | let | 
| 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 295 | val heuristic = Config.get ctxt e_selection_heuristic | 
| 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 296 | val ehoh = string_ord (getenv "E_VERSION", "2.3") <> LESS | 
| 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 297 | val (format, enc) = | 
| 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 298 | if ehoh then (THF (Monomorphic, THF_Predicate_Free), "mono_native_higher") | 
| 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 299 | else (TFF Monomorphic, "mono_native") | 
| 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 300 | in | 
| 43474 
423cd1ecf714
optimized E's time slicing, based on latest exhaustive Judgment Day results
 blanchet parents: 
43473diff
changeset | 301 | (* FUDGE *) | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 302 | if heuristic = e_smartN then | 
| 70939 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 303 | [(0.15, (((128, meshN), format, enc, combsN, false), e_fun_weightN)), | 
| 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 304 | (0.15, (((128, mashN), format, enc, combsN, false), e_sym_offset_weightN)), | 
| 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 305 | (0.15, (((91, mepoN), format, enc, combsN, false), e_autoN)), | 
| 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 306 | (0.15, (((1000, meshN), format, "poly_guards??", combsN, false), e_sym_offset_weightN)), | 
| 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 307 | (0.15, (((256, mepoN), format, enc, liftingN, false), e_fun_weightN)), | 
| 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 308 | (0.25, (((64, mashN), format, enc, combsN, false), e_fun_weightN))] | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 309 | else | 
| 70939 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 310 | [(1.0, (((500, ""), format, enc, combsN, false), heuristic))] | 
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 311 | end, | 
| 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 312 | best_max_mono_iters = default_max_mono_iters, | 
| 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 313 | best_max_new_mono_instances = default_max_new_mono_instances} | 
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 314 | |
| 47646 | 315 | val e = (eN, fn () => e_config) | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 316 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 317 | |
| 70938 | 318 | (* E-Par *) | 
| 48651 | 319 | |
| 70938 | 320 | val e_par_config : atp_config = | 
| 321 |   {exec = K (["E_HOME"], ["runepar.pl"]),
 | |
| 50927 | 322 | arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => | 
| 70938 | 323 | string_of_int (to_secs 1 timeout) ^ " 1 " (* SInE *) ^ file_name ^ " 2" (* proofs *), | 
| 48651 | 324 | proof_delims = tstp_proof_delims, | 
| 325 | known_failures = #known_failures e_config, | |
| 326 | prem_role = Conjecture, | |
| 327 | best_slices = | |
| 328 | (* FUDGE *) | |
| 51018 | 329 | K [(0.25, (((500, meshN), FOF, "mono_guards??", combs_or_liftingN, false), "")), | 
| 330 | (0.25, (((150, meshN), FOF, "poly_tags??", combs_or_liftingN, false), "")), | |
| 331 | (0.25, (((50, meshN), FOF, "mono_tags??", combs_or_liftingN, false), "")), | |
| 332 | (0.25, (((1000, meshN), FOF, "poly_guards??", combsN, false), ""))], | |
| 48651 | 333 | best_max_mono_iters = default_max_mono_iters, | 
| 334 | best_max_new_mono_instances = default_max_new_mono_instances} | |
| 335 | ||
| 50927 | 336 | val e_par = (e_parN, fn () => e_par_config) | 
| 337 | ||
| 338 | ||
| 48700 | 339 | (* iProver *) | 
| 340 | ||
| 341 | val iprover_config : atp_config = | |
| 70932 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 342 |   {exec = K (["IPROVER_HOME"], ["iproveropt", "iprover"]),
 | 
| 50927 | 343 | arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => | 
| 70932 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 344 | "--clausifier \"$E_HOME\"/eprover " ^ | 
| 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 345 | "--clausifier_options \"--tstp-format --silent --cnf\" " ^ | 
| 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 346 | "--time_out_real " ^ string_of_real (Time.toReal timeout) ^ " " ^ file_name, | 
| 48700 | 347 | proof_delims = tstp_proof_delims, | 
| 348 | known_failures = | |
| 349 | [(ProofIncomplete, "% SZS output start CNFRefutation")] @ | |
| 350 | known_szs_status_failures, | |
| 351 | prem_role = Hypothesis, | |
| 352 | best_slices = | |
| 353 | (* FUDGE *) | |
| 51011 | 354 | K [(1.0, (((150, ""), FOF, "mono_guards??", liftingN, false), ""))], | 
| 48700 | 355 | best_max_mono_iters = default_max_mono_iters, | 
| 356 | best_max_new_mono_instances = default_max_new_mono_instances} | |
| 357 | ||
| 358 | val iprover = (iproverN, fn () => iprover_config) | |
| 359 | ||
| 360 | ||
| 44099 | 361 | (* LEO-II *) | 
| 362 | ||
| 363 | val leo2_config : atp_config = | |
| 52754 
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
 blanchet parents: 
52151diff
changeset | 364 |   {exec = K (["LEO2_HOME"], ["leo.opt", "leo"]),
 | 
| 57759 | 365 | arguments = fn _ => fn full_proofs => fn _ => fn timeout => fn file_name => fn _ => | 
| 366 | "--foatp e --atp e=\"$E_HOME\"/eprover \ | |
| 367 | \--atp epclextract=\"$E_HOME\"/epclextract \ | |
| 368 | \--proofoutput 1 --timeout " ^ string_of_int (to_secs 1 timeout) ^ " " ^ | |
| 369 | (if full_proofs then "--notReplLeibnizEQ --notReplAndrewsEQ --notUseExtCnfCmbd " else "") ^ | |
| 370 | file_name, | |
| 44099 | 371 | proof_delims = tstp_proof_delims, | 
| 45207 | 372 | known_failures = | 
| 47974 
08d2dcc2dab9
improved LEO-II definition handling -- still hoping for a fix directly in LEO-II
 blanchet parents: 
47972diff
changeset | 373 | [(TimedOut, "CPU time limit exceeded, terminating"), | 
| 47972 | 374 | (GaveUp, "No.of.Axioms")] @ | 
| 375 | known_szs_status_failures, | |
| 47976 | 376 | prem_role = Hypothesis, | 
| 47914 
94f37848b7c9
LEO-II's "--sos" option confusingly disables rather than enables SOS, and SOS seems to be ignored anyway; also, pass a number of facts that's more appropriate for each prover
 blanchet parents: 
47912diff
changeset | 377 | best_slices = | 
| 44099 | 378 | (* FUDGE *) | 
| 70932 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 379 | K [(1.0, (((40, ""), THF (Monomorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN, false), ""))], | 
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 380 | best_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *), | 
| 53515 
f5b678b155f6
adjusted number of generated monomorphic instances for new monomorphizer based on new evaluation (E, SPASS, Vampire)
 blanchet parents: 
53225diff
changeset | 381 | best_max_new_mono_instances = default_max_new_mono_instances} | 
| 39491 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39375diff
changeset | 382 | |
| 47646 | 383 | val leo2 = (leo2N, fn () => leo2_config) | 
| 44099 | 384 | |
| 385 | ||
| 67021 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 386 | (* Leo-III *) | 
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 387 | |
| 69717 | 388 | (* 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: 
66544diff
changeset | 389 | val leo3_config : atp_config = | 
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 390 |   {exec = K (["LEO3_HOME"], ["leo3"]),
 | 
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 391 | arguments = fn _ => fn full_proofs => fn _ => fn timeout => fn file_name => fn _ => | 
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 392 | file_name ^ " " ^ "--atp cvc=$CVC4_SOLVER --atp e=\"$E_HOME\"/eprover \ | 
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 393 | \-p -t " ^ string_of_int (to_secs 1 timeout) ^ " " ^ | 
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 394 | (if full_proofs then "--nleq --naeq " else ""), | 
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 395 | proof_delims = tstp_proof_delims, | 
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 396 | known_failures = known_szs_status_failures, | 
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 397 | prem_role = Hypothesis, | 
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 398 | best_slices = | 
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 399 | (* FUDGE *) | 
| 70940 
ce3a05ad07b7
added support for Zipperposition on SystemOnTPTP
 blanchet parents: 
70939diff
changeset | 400 | K [(1.0, (((150, ""), THF (Polymorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN, false), ""))], | 
| 67021 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 401 | best_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *), | 
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 402 | best_max_new_mono_instances = default_max_new_mono_instances} | 
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 403 | |
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 404 | val leo3 = (leo3N, fn () => leo3_config) | 
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 405 | |
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 406 | |
| 44099 | 407 | (* Satallax *) | 
| 408 | ||
| 52097 | 409 | (* Choice is disabled until there is proper reconstruction for it. *) | 
| 44099 | 410 | val satallax_config : atp_config = | 
| 52754 
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
 blanchet parents: 
52151diff
changeset | 411 |   {exec = K (["SATALLAX_HOME"], ["satallax.opt", "satallax"]),
 | 
| 50927 | 412 | arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => | 
| 70929 | 413 | (case getenv "E_HOME" of | 
| 414 | "" => "" | |
| 415 | | home => "-E " ^ home ^ "/eprover ") ^ | |
| 68563 | 416 | "-p tstp -t " ^ string_of_int (to_secs 1 timeout) ^ " " ^ file_name, | 
| 45162 | 417 | proof_delims = | 
| 57707 
0242e9578828
imported patch satallax_proof_support_Sledgehammer
 fleury parents: 
57672diff
changeset | 418 |      [("% SZS output start Proof", "% SZS output end Proof")],
 | 
| 45203 | 419 | known_failures = known_szs_status_failures, | 
| 47981 | 420 | prem_role = Hypothesis, | 
| 44416 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 blanchet parents: 
44391diff
changeset | 421 | best_slices = | 
| 44754 | 422 | (* FUDGE *) | 
| 70932 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 423 | K [(1.0, (((150, ""), THF (Monomorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN, false), ""))], | 
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 424 | best_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *), | 
| 53515 
f5b678b155f6
adjusted number of generated monomorphic instances for new monomorphizer based on new evaluation (E, SPASS, Vampire)
 blanchet parents: 
53225diff
changeset | 425 | best_max_new_mono_instances = default_max_new_mono_instances} | 
| 44099 | 426 | |
| 47646 | 427 | val satallax = (satallaxN, fn () => satallax_config) | 
| 44099 | 428 | |
| 429 | ||
| 430 | (* 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: 
42723diff
changeset | 431 | |
| 48005 
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
 blanchet parents: 
48004diff
changeset | 432 | 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: 
49991diff
changeset | 433 | 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: 
48004diff
changeset | 434 | 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: 
48004diff
changeset | 435 | 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: 
48004diff
changeset | 436 | 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: 
49991diff
changeset | 437 | 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: 
47053diff
changeset | 438 | |
| 48005 
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
 blanchet parents: 
48004diff
changeset | 439 | val spass_config : atp_config = | 
| 66363 | 440 |   {exec = K (["SPASS_HOME"], ["SPASS"]),
 | 
| 68563 | 441 | arguments = fn _ => fn full_proofs => fn extra_options => fn timeout => fn file_name => fn _ => | 
| 442 | "-Isabelle=1 " ^ (if full_proofs then "-CNFRenaming=0 -Splits=0 " else "") ^ | |
| 443 | "-TimeLimit=" ^ string_of_int (to_secs 1 timeout) ^ " " ^ file_name | |
| 444 | |> extra_options <> "" ? prefix (extra_options ^ " "), | |
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 445 |    proof_delims = [("Here is a proof", "Formulae used in the proof")],
 | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 446 | known_failures = | 
| 53225 
16235bb41881
got rid of old error -- users who install SPASS manually are responsible for any version mismatches
 blanchet parents: 
52995diff
changeset | 447 | [(GaveUp, "SPASS beiseite: Completion found"), | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 448 | (TimedOut, "SPASS beiseite: Ran out of time"), | 
| 36965 | 449 | (OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"), | 
| 37413 | 450 | (MalformedInput, "Undefined symbol"), | 
| 37414 
d0cea0796295
expect SPASS 3.7, and give a friendly warning if an older version is used
 blanchet parents: 
37413diff
changeset | 451 | (MalformedInput, "Free Variable"), | 
| 44391 | 452 | (Unprovable, "No formulae and clauses found in input file"), | 
| 47972 | 453 | (InternalError, "Please report this error")] @ | 
| 454 | known_perl_failures, | |
| 47976 | 455 | prem_role = Conjecture, | 
| 71793 | 456 | best_slices = fn _ => | 
| 42723 | 457 | (* FUDGE *) | 
| 51016 | 458 | [(0.1667, (((150, meshN), DFG Monomorphic, "mono_native", combsN, true), "")), | 
| 459 | (0.1667, (((500, meshN), DFG Monomorphic, "mono_native", liftingN, true), spass_H2SOS)), | |
| 460 | (0.1666, (((50, meshN), DFG Monomorphic, "mono_native", liftingN, true), spass_H2LR0LT0)), | |
| 461 | (0.1000, (((250, meshN), DFG Monomorphic, "mono_native", combsN, true), spass_H2NuVS0)), | |
| 462 | (0.1000, (((1000, mepoN), DFG Monomorphic, "mono_native", liftingN, true), spass_H1SOS)), | |
| 463 | (0.1000, (((150, meshN), DFG Monomorphic, "poly_guards??", liftingN, false), spass_H2NuVS0Red2)), | |
| 464 | (0.1000, (((300, meshN), DFG Monomorphic, "mono_native", combsN, true), spass_H2SOS)), | |
| 71353 | 465 | (0.1000, (((100, meshN), DFG Monomorphic, "mono_native", combs_and_liftingN, true), spass_H2))], | 
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 466 | best_max_mono_iters = default_max_mono_iters, | 
| 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 467 | best_max_new_mono_instances = default_max_new_mono_instances} | 
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 468 | |
| 48005 
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
 blanchet parents: 
48004diff
changeset | 469 | val spass = (spassN, fn () => spass_config) | 
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 470 | |
| 52073 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 471 | |
| 37509 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 472 | (* Vampire *) | 
| 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 473 | |
| 68563 | 474 | fun is_vampire_noncommercial_license_accepted () = | 
| 475 | let | |
| 69593 | 476 | val flag = Options.default_string \<^system_option>\<open>vampire_noncommercial\<close> | 
| 68563 | 477 | |> String.map Char.toLower | 
| 478 | in | |
| 479 | if flag = "yes" then | |
| 480 | SOME true | |
| 481 | else if flag = "no" then | |
| 482 | SOME false | |
| 483 | else | |
| 484 | NONE | |
| 485 | end | |
| 486 | ||
| 487 | fun check_vampire_noncommercial () = | |
| 488 | (case is_vampire_noncommercial_license_accepted () of | |
| 489 | SOME true => () | |
| 490 | | SOME false => | |
| 491 | error (Pretty.string_of (Pretty.para | |
| 492 | "The Vampire prover may be used only for noncommercial applications")) | |
| 493 | | NONE => | |
| 494 | error (Pretty.string_of (Pretty.para | |
| 495 | "The Vampire prover is not activated; to activate it, set the Isabelle system option \ | |
| 496 | \\"vampire_noncommercial\" to \"yes\" (e.g. via the Isabelle/jEdit menu Plugin Options \ | |
| 497 | \/ Isabelle / General)"))) | |
| 44420 | 498 | |
| 68563 | 499 | val vampire_basic_options = "--proof tptp --output_axiom_names on $VAMPIRE_EXTRA_OPTIONS" | 
| 58084 | 500 | |
| 501 | val vampire_full_proof_options = | |
| 71793 | 502 | " --proof_extra free --forced_options avatar=off:equality_proxy=off:general_splitting=off:inequality_splitting=0:naming=0" | 
| 58084 | 503 | |
| 70941 | 504 | val remote_vampire_command = | 
| 58084 | 505 | "vampire " ^ vampire_basic_options ^ " " ^ vampire_full_proof_options ^ " -t %d %s" | 
| 506 | ||
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 507 | val vampire_config : atp_config = | 
| 52754 
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
 blanchet parents: 
52151diff
changeset | 508 |   {exec = K (["VAMPIRE_HOME"], ["vampire"]),
 | 
| 58084 | 509 | arguments = fn _ => fn full_proofs => fn sos => fn timeout => fn file_name => fn _ => | 
| 68563 | 510 | (check_vampire_noncommercial (); | 
| 511 | vampire_basic_options ^ (if full_proofs then " " ^ vampire_full_proof_options else "") ^ | |
| 512 | " -t " ^ string_of_int (to_secs 1 timeout) ^ " --input_file " ^ file_name | |
| 513 | |> sos = sosN ? prefix "--sos on "), | |
| 37509 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 514 | proof_delims = | 
| 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 515 |      [("=========== Refutation ==========",
 | 
| 52073 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 516 | "======= End of refutation =======")] @ | 
| 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 517 | tstp_proof_delims, | 
| 37509 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 518 | known_failures = | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
42999diff
changeset | 519 | [(GaveUp, "UNPROVABLE"), | 
| 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
42999diff
changeset | 520 | (GaveUp, "CANNOT PROVE"), | 
| 37509 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 521 | (Unprovable, "Satisfiability detected"), | 
| 38647 
5500241da479
play with fudge factor + parse one more Vampire error
 blanchet parents: 
38646diff
changeset | 522 | (Unprovable, "Termination reason: Satisfiable"), | 
| 47972 | 523 | (Interrupted, "Aborted by signal SIGINT")] @ | 
| 524 | known_szs_status_failures, | |
| 63116 | 525 | prem_role = Hypothesis, | 
| 42725 
64dea91bbe0e
added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
 blanchet parents: 
42723diff
changeset | 526 | best_slices = fn ctxt => | 
| 42723 | 527 | (* FUDGE *) | 
| 70932 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 528 | [(0.333, (((500, meshN), TFF Monomorphic, "mono_native", combs_or_liftingN, false), sosN)), | 
| 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 529 | (0.333, (((150, meshN), TFF Monomorphic, "poly_tags??", combs_or_liftingN, false), sosN)), | 
| 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 530 | (0.334, (((50, meshN), TFF Monomorphic, "mono_native", combs_or_liftingN, false), no_sosN))] | 
| 58084 | 531 | |> Config.get ctxt force_sos ? (hd #> apfst (K 1.0) #> single), | 
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 532 | best_max_mono_iters = default_max_mono_iters, | 
| 53515 
f5b678b155f6
adjusted number of generated monomorphic instances for new monomorphizer based on new evaluation (E, SPASS, Vampire)
 blanchet parents: 
53225diff
changeset | 533 | best_max_new_mono_instances = 2 * default_max_new_mono_instances (* FUDGE *)} | 
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 534 | |
| 47646 | 535 | val vampire = (vampireN, fn () => vampire_config) | 
| 37509 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 536 | |
| 68563 | 537 | |
| 48803 
ffa31bf5c662
tone down "z3_tptp", now that Z3 (starting with 4.1) no longer supports TPTP TFF0
 blanchet parents: 
48801diff
changeset | 538 | (* Z3 with TPTP syntax (half experimental, half legacy) *) | 
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 539 | |
| 44423 
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
 blanchet parents: 
44422diff
changeset | 540 | val z3_tptp_config : atp_config = | 
| 56378 | 541 |   {exec = K (["Z3_TPTP_HOME"], ["z3_tptp"]),
 | 
| 50927 | 542 | arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => | 
| 57261 | 543 | "-proof -t:" ^ string_of_int (to_secs 1 timeout) ^ " -file:" ^ file_name, | 
| 56397 | 544 |    proof_delims = [("SZS status Theorem", "")],
 | 
| 45203 | 545 | known_failures = known_szs_status_failures, | 
| 47976 | 546 | prem_role = Hypothesis, | 
| 42723 | 547 | best_slices = | 
| 44423 
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
 blanchet parents: 
44422diff
changeset | 548 | (* FUDGE *) | 
| 70932 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 549 | K [(0.5, (((250, meshN), TFF Monomorphic, "mono_native", combsN, false), "")), | 
| 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 550 | (0.25, (((125, mepoN), TFF Monomorphic, "mono_native", combsN, false), "")), | 
| 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 551 | (0.125, (((62, mashN), TFF Monomorphic, "mono_native", combsN, false), "")), | 
| 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 552 | (0.125, (((31, meshN), TFF Monomorphic, "mono_native", combsN, false), ""))], | 
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 553 | best_max_mono_iters = default_max_mono_iters, | 
| 53515 
f5b678b155f6
adjusted number of generated monomorphic instances for new monomorphizer based on new evaluation (E, SPASS, Vampire)
 blanchet parents: 
53225diff
changeset | 554 | best_max_new_mono_instances = 2 * default_max_new_mono_instances (* FUDGE *)} | 
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 555 | |
| 47646 | 556 | val z3_tptp = (z3_tptpN, fn () => z3_tptp_config) | 
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 557 | |
| 44590 | 558 | |
| 69717 | 559 | (* Zipperposition *) | 
| 560 | ||
| 72174 | 561 | val zipperposition_blsimp = "--mode ho-pragmatic --max-inferences 3 --ho-max-app-projections 0 --ho-max-elims 0 --ho-max-rigid-imitations 2 --ho-max-identifications 0 --ho-unif-max-depth 2 --boolean-reasoning no-cases --ext-rules ext-family --ext-rules-max-depth 1 --kbo-weight-fun invdocc --ho-prim-enum tf --ho-prim-enum-early-bird true --tptp-def-as-rewrite --ho-unif-level pragmatic-framework -q '1|const|conjecture-relative-var(1,s,f)' -q '1|prefer-processed|pnrefined(1,1,1,2,2,2,0.5)' -q '1|prefer-sos|staggered(1)' -q '2|prefer-fo|default' -q '1|prefer-neg-unit|orient-lmax(2,1,2,1,1)' -q '2|prefer-easy-ho|conjecture-relative-struct(1.5,3.5,2,3)' --ho-elim-leibniz 2 --ho-fixpoint-decider true --ho-pattern-decider false --ho-solid-decider true --ho-max-solidification 12 --select e-selection11 --solve-formulas true --sup-at-vars false --sup-at-var-headed false --lazy-cnf true --lazy-cnf-kind simp --lazy-cnf-renaming-threshold 4 --sine 50 --sine-tolerance 1.7 --sine-depth-max 3 --sine-depth-min 1 --sine-trim-implications true --ho-selection-restriction none --sup-from-var-headed false --sine-trim-implications true" | 
| 562 | val zipperposition_s6 = "--tptp-def-as-rewrite --rewrite-before-cnf true --mode ho-competitive --boolean-reasoning no-cases --ext-rules off --ho-prim-enum none --recognize-injectivity true --ho-elim-leibniz off --ho-unif-level full-framework --no-max-vars -q '3|const|conjecture-relative-var(1.02,l,f)' -q '1|prefer-ho-steps|conjecture-relative-var(1,s,f)' -q '1|prefer-processed|fifo' -q '3|by-app-var-num|pnrefined(2,1,1,1,2,2,2)' --select ho-selection5 --prec-gen-fun unary_first --solid-subsumption false --ignore-orphans false --ho-solid-decider true --ho-fixpoint-decider true --ho-pattern-decider true --sup-at-vars false --sup-at-var-headed false --sup-from-var-headed false --ho-neg-ext-simpl true" | |
| 563 | val zipperposition_cdots = "--mode ho-competitive --boolean-reasoning cases-simpl --ext-rules ext-family --ext-rules-max-depth 1 --ho-prim-enum pragmatic --ho-prim-max 1 --bool-subterm-selection A --avatar off --recognize-injectivity true --ho-elim-leibniz 1 --ho-unif-level full-framework --no-max-vars -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 --sine 50 --sine-tolerance 2 --sine-depth-max 4 --sine-depth-min 1 --scan-clause-ac true --lambdasup 0 --kbo-weight-fun invfreqrank" | |
| 564 | ||
| 57154 | 565 | val zipperposition_config : atp_config = | 
| 566 |   {exec = K (["ZIPPERPOSITION_HOME"], ["zipperposition"]),
 | |
| 72174 | 567 | arguments = fn _ => fn _ => fn extra_options => fn timeout => fn file_name => fn _ => | 
| 568 | "--input tptp --output tptp --timeout " ^ string_of_int (to_secs 1 timeout) ^ " " ^ file_name | |
| 569 | |> extra_options <> "" ? prefix (extra_options ^ " "), | |
| 57154 | 570 | proof_delims = tstp_proof_delims, | 
| 571 | known_failures = known_szs_status_failures, | |
| 72174 | 572 | prem_role = Conjecture, | 
| 57154 | 573 | best_slices = fn _ => | 
| 574 | (* FUDGE *) | |
| 72174 | 575 | [(0.333, (((128, "meshN"), THF (Monomorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN, false), zipperposition_blsimp)), | 
| 576 | (0.333, (((32, "meshN"), THF (Polymorphic, THF_Without_Choice), "poly_native_higher", keep_lamsN, false), zipperposition_s6)), | |
| 577 | (0.334, (((512, "meshN"), THF (Monomorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN, false), zipperposition_cdots))], | |
| 57154 | 578 | best_max_mono_iters = default_max_mono_iters, | 
| 579 | best_max_new_mono_instances = default_max_new_mono_instances} | |
| 580 | ||
| 581 | val zipperposition = (zipperpositionN, fn () => zipperposition_config) | |
| 582 | ||
| 583 | ||
| 70932 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 584 | (* Remote Pirate invocation *) | 
| 44590 | 585 | |
| 59577 | 586 | val remote_pirate_config : atp_config = | 
| 587 |   {exec = K (["ISABELLE_ATP"], ["scripts/remote_pirate"]),
 | |
| 54788 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 588 | arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => | 
| 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 589 | string_of_int (to_secs 1 timeout) ^ " " ^ file_name, | 
| 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 590 |    proof_delims = [("Involved clauses:", "Involved clauses:")],
 | 
| 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 591 | known_failures = known_szs_status_failures, | 
| 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 592 | prem_role = #prem_role spass_config, | 
| 70932 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 593 | best_slices = K [(1.0, (((200, ""), DFG Polymorphic, "tc_native", combsN, true), ""))], | 
| 54788 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 594 | best_max_mono_iters = default_max_mono_iters, | 
| 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 595 | best_max_new_mono_instances = default_max_new_mono_instances} | 
| 59577 | 596 | val remote_pirate = (remote_prefix ^ pirateN, fn () => remote_pirate_config) | 
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 597 | |
| 52073 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 598 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 599 | (* Remote ATP invocation via SystemOnTPTP *) | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 600 | |
| 49984 | 601 | val remote_systems = Synchronized.var "atp_remote_systems" ([] : string list) | 
| 31835 | 602 | |
| 49984 | 603 | fun get_remote_systems () = | 
| 62519 | 604 | Timeout.apply (seconds 10.0) (fn () => | 
| 58084 | 605 | (case Isabelle_System.bash_output "\"$ISABELLE_ATP/scripts/remote_atp\" -w 2>&1" of | 
| 606 | (output, 0) => split_lines output | |
| 607 | | (output, _) => | |
| 608 | (warning | |
| 609 | (case extract_known_atp_failure known_perl_failures output of | |
| 610 | SOME failure => string_of_atp_failure failure | |
| 63692 | 611 | | NONE => trim_line output); []))) () | 
| 62519 | 612 | handle Timeout.TIMEOUT _ => [] | 
| 31835 | 613 | |
| 49984 | 614 | 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: 
42535diff
changeset | 615 | find_first (String.isPrefix (name ^ "---")) systems | 
| 49984 | 616 | | 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: 
38685diff
changeset | 617 | case find_first (String.isPrefix (name ^ "---" ^ version)) systems of | 
| 49984 | 618 | 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: 
38685diff
changeset | 619 | | res => res | 
| 
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
 blanchet parents: 
38685diff
changeset | 620 | |
| 49984 | 621 | fun get_remote_system name versions = | 
| 70942 
718255bde391
invoke remote Vampire with higher-order (THF) syntax
 blanchet parents: 
70941diff
changeset | 622 | Synchronized.change_result remote_systems (fn systems => | 
| 
718255bde391
invoke remote Vampire with higher-order (THF) syntax
 blanchet parents: 
70941diff
changeset | 623 | (if null systems then get_remote_systems () else systems) | 
| 
718255bde391
invoke remote Vampire with higher-order (THF) syntax
 blanchet parents: 
70941diff
changeset | 624 | |> `(`(find_remote_system name versions))) | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 625 | |
| 49984 | 626 | fun the_remote_system name versions = | 
| 54788 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 627 | (case get_remote_system name versions of | 
| 42955 | 628 | (SOME sys, _) => sys | 
| 63692 | 629 | | (NONE, []) => error "SystemOnTPTP is currently not available" | 
| 42955 | 630 | | (NONE, syss) => | 
| 54788 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 631 | (case syss |> filter_out (String.isPrefix "%") |> filter_out (curry (op =) "") of | 
| 63692 | 632 | [] => error "SystemOnTPTP is currently not available" | 
| 633 |     | [msg] => error ("SystemOnTPTP is currently not available: " ^ msg)
 | |
| 46480 | 634 | | syss => | 
| 54788 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 635 |       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: 
54197diff
changeset | 636 | commas_quote syss ^ ".)"))) | 
| 31835 | 637 | |
| 72174 | 638 | val max_remote_secs = 1000 (* give Geoff Sutcliffe's servers a break *) | 
| 41148 | 639 | |
| 58084 | 640 | fun remote_config system_name system_versions proof_delims known_failures prem_role best_slice = | 
| 52754 
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
 blanchet parents: 
52151diff
changeset | 641 |   {exec = K (["ISABELLE_ATP"], ["scripts/remote_atp"]),
 | 
| 70941 | 642 | arguments = fn _ => fn _ => fn command => fn timeout => fn file_name => fn _ => | 
| 643 | (if command <> "" then "-c " ^ quote command ^ " " else "") ^ | |
| 54788 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 644 | "-s " ^ the_remote_system system_name system_versions ^ " " ^ | 
| 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 645 | "-t " ^ string_of_int (Int.min (max_remote_secs, to_secs 1 timeout)) ^ | 
| 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 646 | " " ^ file_name, | 
| 42962 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 blanchet parents: 
42955diff
changeset | 647 | proof_delims = union (op =) tstp_proof_delims proof_delims, | 
| 45203 | 648 | known_failures = known_failures @ known_perl_failures @ known_says_failures, | 
| 47976 | 649 | prem_role = prem_role, | 
| 48716 
1d2a12bb0640
stop distinguishing between complete and incomplete slices, since this is very fragile and has hardly any useful semantics to users
 blanchet parents: 
48715diff
changeset | 650 | best_slices = fn ctxt => [(1.0, best_slice ctxt)], | 
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47955diff
changeset | 651 | best_max_mono_iters = default_max_mono_iters, | 
| 58084 | 652 | best_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: 
42332diff
changeset | 653 | |
| 43500 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 654 | fun remotify_config system_name system_versions best_slice | 
| 58084 | 655 |     ({proof_delims, known_failures, prem_role, ...} : atp_config) =
 | 
| 656 | remote_config system_name system_versions proof_delims known_failures prem_role best_slice | |
| 38023 | 657 | |
| 58084 | 658 | fun remote_atp name system_name system_versions proof_delims known_failures prem_role best_slice = | 
| 659 | (remote_prefix ^ name, fn () => | |
| 660 | remote_config system_name system_versions proof_delims known_failures prem_role best_slice) | |
| 43500 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 661 | fun remotify_atp (name, config) system_name system_versions best_slice = | 
| 58084 | 662 | (remote_prefix ^ name, remotify_config system_name system_versions best_slice o config) | 
| 28592 | 663 | |
| 57269 
1df6f747f164
changed type encoding for new Waldmeister, to trigger filtering of 'dangerous' lemmas
 blanchet parents: 
57265diff
changeset | 664 | fun gen_remote_waldmeister name type_enc = | 
| 57265 | 665 | remote_atp name "Waldmeister" ["710"] tstp_proof_delims | 
| 666 | ([(OutOfResources, "Too many function symbols"), | |
| 667 | (Inappropriate, "**** Unexpected end of file."), | |
| 668 | (Crashed, "Unrecoverable Segmentation Fault")] | |
| 669 | @ known_szs_status_failures) | |
| 57264 | 670 | Hypothesis | 
| 57269 
1df6f747f164
changed type encoding for new Waldmeister, to trigger filtering of 'dangerous' lemmas
 blanchet parents: 
57265diff
changeset | 671 | (K (((50, ""), CNF_UEQ, type_enc, combsN, false), "") (* FUDGE *)) | 
| 57264 | 672 | |
| 52094 | 673 | val remote_agsyhol = | 
| 674 | remotify_atp agsyhol "agsyHOL" ["1.0", "1"] | |
| 70932 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 675 | (K (((60, ""), THF (Monomorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN, false), "") (* FUDGE *)) | 
| 70937 | 676 | val remote_alt_ergo = | 
| 677 | remotify_atp alt_ergo "Alt-Ergo" ["0.95.2"] | |
| 678 | (K (((250, ""), TFF Polymorphic, "poly_native", keep_lamsN, false), "") (* FUDGE *)) | |
| 43500 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 679 | val remote_e = | 
| 63768 | 680 | remotify_atp e "E" ["2.0", "1.9.1", "1.8"] | 
| 70932 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 681 | (K (((750, ""), TFF Monomorphic, "mono_native", combsN, false), "") (* FUDGE *)) | 
| 48700 | 682 | val remote_iprover = | 
| 52094 | 683 | remotify_atp iprover "iProver" ["0.99"] | 
| 58084 | 684 | (K (((150, ""), FOF, "mono_guards??", liftingN, false), "") (* FUDGE *)) | 
| 44099 | 685 | val remote_leo2 = | 
| 52094 | 686 | remotify_atp leo2 "LEO-II" ["1.5.0", "1.4", "1.3", "1.2", "1"] | 
| 70932 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 687 | (K (((40, ""), THF (Monomorphic, 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: 
66544diff
changeset | 688 | val remote_leo3 = | 
| 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 blanchet parents: 
66544diff
changeset | 689 | remotify_atp leo3 "Leo-III" ["1.1"] | 
| 70932 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 690 | (K (((150, ""), THF (Polymorphic, THF_Without_Choice), "poly_native_higher", keep_lamsN, false), "") (* FUDGE *)) | 
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 691 | val remote_snark = | 
| 52094 | 692 | remote_atp snarkN "SNARK" ["20120808r022", "20080805r029", "20080805r024"] | 
| 58084 | 693 |     [("refutation.", "end_refutation.")] [] Hypothesis
 | 
| 70932 
a35618d00d29
updated iProver setup and tuned other ATP setups
 blanchet parents: 
70931diff
changeset | 694 | (K (((100, ""), TFF Monomorphic, "mono_native", liftingN, false), "") (* FUDGE *)) | 
| 70937 | 695 | val remote_vampire = | 
| 70942 
718255bde391
invoke remote Vampire with higher-order (THF) syntax
 blanchet parents: 
70941diff
changeset | 696 | remotify_atp vampire "Vampire" ["THF-4.4"] | 
| 
718255bde391
invoke remote Vampire with higher-order (THF) syntax
 blanchet parents: 
70941diff
changeset | 697 | (K (((400, ""), THF (Monomorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN, false), remote_vampire_command) (* FUDGE *)) | 
| 57269 
1df6f747f164
changed type encoding for new Waldmeister, to trigger filtering of 'dangerous' lemmas
 blanchet parents: 
57265diff
changeset | 698 | val remote_waldmeister = gen_remote_waldmeister waldmeisterN "raw_mono_tags??" | 
| 70940 
ce3a05ad07b7
added support for Zipperposition on SystemOnTPTP
 blanchet parents: 
70939diff
changeset | 699 | val remote_zipperposition = | 
| 72174 | 700 | remotify_atp zipperposition "Zipperpin" ["2.0"] | 
| 701 | (K (((512, ""), THF (Monomorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN, false), "") (* FUDGE *)) | |
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 702 | |
| 52073 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 703 | |
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 704 | (* Setup *) | 
| 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 705 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 706 | fun add_atp (name, config) thy = | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 707 | Data.map (Symtab.update_new (name, (config, stamp ()))) thy | 
| 63692 | 708 |   handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name)
 | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 709 | |
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 710 | fun get_atp thy name = | 
| 54788 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 711 | fst (the (Symtab.lookup (Data.get thy) name)) | 
| 63692 | 712 |   handle Option.Option => error ("Unknown ATP: " ^ name)
 | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 713 | |
| 41727 
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
 blanchet parents: 
41725diff
changeset | 714 | val supported_atps = Symtab.keys o Data.get | 
| 36371 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 715 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 716 | fun is_atp_installed thy name = | 
| 48376 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48232diff
changeset | 717 |   let val {exec, ...} = get_atp thy name () in
 | 
| 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: 
57636diff
changeset | 718 | exists (fn var => getenv var <> "") (fst (exec false)) | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 719 | end | 
| 36371 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 720 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 721 | fun refresh_systems_on_tptp () = | 
| 49984 | 722 | Synchronized.change remote_systems (fn _ => get_remote_systems ()) | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 723 | |
| 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: 
47053diff
changeset | 724 | fun effective_term_order ctxt atp = | 
| 
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
 blanchet parents: 
47053diff
changeset | 725 | let val ord = Config.get ctxt term_order in | 
| 
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
 blanchet parents: 
47053diff
changeset | 726 | if ord = smartN then | 
| 54788 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 727 |       {is_lpo = false, gen_weights = (atp = spassN), gen_prec = (atp = spassN),
 | 
| 59577 | 728 | gen_simp = String.isSuffix pirateN atp} | 
| 47055 
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
 blanchet parents: 
47053diff
changeset | 729 | else | 
| 
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
 blanchet parents: 
47053diff
changeset | 730 | let val is_lpo = String.isSubstring lpoN ord in | 
| 54788 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 731 |         {is_lpo = is_lpo, gen_weights = not is_lpo andalso String.isSubstring xweightsN ord,
 | 
| 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 732 | gen_prec = String.isSubstring xprecN ord, gen_simp = String.isSubstring xsimpN ord} | 
| 47055 
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
 blanchet parents: 
47053diff
changeset | 733 | end | 
| 
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
 blanchet parents: 
47053diff
changeset | 734 | end | 
| 
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
 blanchet parents: 
47053diff
changeset | 735 | |
| 52073 
ccb292952774
started adding agsyHOL as an experimental prover
 blanchet parents: 
51998diff
changeset | 736 | val atps = | 
| 70939 
3218999b3715
folded experimental Ehoh into E now that E 2.3 has been released
 blanchet parents: 
70938diff
changeset | 737 | [agsyhol, alt_ergo, e, e_par, iprover, leo2, leo3, satallax, spass, vampire, z3_tptp, | 
| 70938 | 738 | zipperposition, remote_agsyhol, remote_alt_ergo, remote_e, remote_iprover, remote_leo2, | 
| 70940 
ce3a05ad07b7
added support for Zipperposition on SystemOnTPTP
 blanchet parents: 
70939diff
changeset | 739 | remote_leo3, remote_pirate, remote_snark, remote_vampire, remote_waldmeister, | 
| 
ce3a05ad07b7
added support for Zipperposition on SystemOnTPTP
 blanchet parents: 
70939diff
changeset | 740 | remote_zipperposition] | 
| 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: 
47053diff
changeset | 741 | |
| 57262 | 742 | val _ = Theory.setup (fold add_atp atps) | 
| 35867 | 743 | |
| 28592 | 744 | end; |