| author | berghofe | 
| Tue, 10 Jan 2012 23:59:37 +0100 | |
| changeset 46181 | 49c3e0ef9d70 | 
| parent 45876 | 40952db4e57b | 
| child 46320 | 0b8b73b49848 | 
| 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 | 
| 45301 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 10 | type atp_format = ATP_Problem.atp_format | 
| 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 | 11 | type formula_kind = ATP_Problem.formula_kind | 
| 39491 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39375diff
changeset | 12 | type failure = ATP_Proof.failure | 
| 38023 | 13 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 14 | type atp_config = | 
| 42578 
1eaf4d437d4c
define type system in ATP module so that ATPs can suggest a type system
 blanchet parents: 
42577diff
changeset | 15 |     {exec : string * string,
 | 
| 
1eaf4d437d4c
define type system in ATP module so that ATPs can suggest a type system
 blanchet parents: 
42577diff
changeset | 16 | required_execs : (string * string) list, | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 17 | arguments : | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 18 | Proof.context -> bool -> string -> Time.time | 
| 43354 
396aaa15dd8b
pass --trim option to "eproof" script to speed up proof reconstruction
 blanchet parents: 
43288diff
changeset | 19 | -> (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, | 
| 
1eaf4d437d4c
define type system in ATP module so that ATPs can suggest a type system
 blanchet parents: 
42577diff
changeset | 21 | known_failures : (failure * string) list, | 
| 42709 
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
 blanchet parents: 
42708diff
changeset | 22 | conj_sym_kind : formula_kind, | 
| 
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
 blanchet parents: 
42708diff
changeset | 23 | prem_kind : formula_kind, | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 24 | best_slices : | 
| 44754 | 25 | Proof.context | 
| 45521 | 26 | -> (real * (bool * (int * atp_format * string * string * string))) list} | 
| 38023 | 27 | |
| 44099 | 28 | val force_sos : bool 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 | 29 | 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 | 30 | 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 | 31 | 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 | 32 | val e_sym_offset_weightN : string | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 33 | val e_weight_method : string Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 34 | val e_default_fun_weight : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 35 | val e_fun_weight_base : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 36 | val e_fun_weight_span : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 37 | val e_default_sym_offs_weight : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 38 | val e_sym_offs_weight_base : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 39 | val e_sym_offs_weight_span : real Config.T | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 40 | val eN : string | 
| 44590 | 41 | val e_sineN : string | 
| 42 | val e_tofofN : string | |
| 45338 | 43 | val iproverN : string | 
| 44 | val iprover_eqN : string | |
| 44590 | 45 | val leo2N : string | 
| 45365 | 46 | val dummy_tff1N : string | 
| 47 | val dummy_thfN : string | |
| 44590 | 48 | val satallaxN : string | 
| 49 | val snarkN : string | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 50 | val spassN : string | 
| 45301 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 51 | val spass_newN : string | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 52 | val vampireN : string | 
| 42938 | 53 | val waldmeisterN : string | 
| 44423 
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
 blanchet parents: 
44422diff
changeset | 54 | val z3_tptpN : string | 
| 40060 
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
 blanchet parents: 
40059diff
changeset | 55 | val remote_prefix : string | 
| 41738 
eb98c60a6cf0
added experimental "remote_z3_atp", Sutcliffe's TPTP-syntax-aware wrapper for Z3 -- allows to do head-to-head comparison of Sledgehammer's ATP translation and of Sascha's SMT translation
 blanchet parents: 
41727diff
changeset | 56 | val remote_atp : | 
| 
eb98c60a6cf0
added experimental "remote_z3_atp", Sutcliffe's TPTP-syntax-aware wrapper for Z3 -- allows to do head-to-head comparison of Sledgehammer's ATP translation and of Sascha's SMT translation
 blanchet parents: 
41727diff
changeset | 57 | string -> string -> string list -> (string * string) list | 
| 44416 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 blanchet parents: 
44391diff
changeset | 58 | -> (failure * string) list -> formula_kind -> formula_kind | 
| 45521 | 59 | -> (Proof.context -> int * atp_format * string * string) | 
| 60 | -> string * atp_config | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 61 | val add_atp : string * atp_config -> theory -> theory | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 62 | val get_atp : theory -> string -> atp_config | 
| 41727 
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
 blanchet parents: 
41725diff
changeset | 63 | val supported_atps : theory -> string list | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 64 | val is_atp_installed : theory -> string -> bool | 
| 35867 | 65 | val refresh_systems_on_tptp : unit -> unit | 
| 66 | val setup : theory -> theory | |
| 28592 | 67 | end; | 
| 68 | ||
| 36376 | 69 | structure ATP_Systems : ATP_SYSTEMS = | 
| 28592 | 70 | struct | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 71 | |
| 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 | 72 | open ATP_Problem | 
| 39491 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39375diff
changeset | 73 | open ATP_Proof | 
| 45521 | 74 | open ATP_Translate | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 75 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 76 | (* ATP configuration *) | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 77 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 78 | type atp_config = | 
| 42578 
1eaf4d437d4c
define type system in ATP module so that ATPs can suggest a type system
 blanchet parents: 
42577diff
changeset | 79 |   {exec : string * string,
 | 
| 
1eaf4d437d4c
define type system in ATP module so that ATPs can suggest a type system
 blanchet parents: 
42577diff
changeset | 80 | required_execs : (string * string) list, | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 81 | arguments : | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 82 | Proof.context -> bool -> string -> Time.time | 
| 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 83 | -> (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 | 84 | proof_delims : (string * string) list, | 
| 
1eaf4d437d4c
define type system in ATP module so that ATPs can suggest a type system
 blanchet parents: 
42577diff
changeset | 85 | known_failures : (failure * string) list, | 
| 42709 
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
 blanchet parents: 
42708diff
changeset | 86 | conj_sym_kind : formula_kind, | 
| 
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
 blanchet parents: 
42708diff
changeset | 87 | prem_kind : formula_kind, | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 88 | best_slices : | 
| 44754 | 89 | Proof.context | 
| 45521 | 90 | -> (real * (bool * (int * atp_format * string * string * string))) list} | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 91 | |
| 42723 | 92 | (* "best_slices" must be found empirically, taking a wholistic approach since | 
| 93 | the ATPs are run in parallel. The "real" components give the faction of the | |
| 43569 
b342cd125533
removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
 blanchet parents: 
43567diff
changeset | 94 | time available given to the slice and should add up to 1.0. The "bool" | 
| 42723 | 95 | component indicates whether the slice's strategy is complete; the "int", the | 
| 43569 
b342cd125533
removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
 blanchet parents: 
43567diff
changeset | 96 | preferred number of facts to pass; the first "string", the preferred type | 
| 45521 | 97 | system (which should be sound or quasi-sound); the second "string", the | 
| 98 | preferred lambda translation scheme; the third "string", extra information to | |
| 99 | the prover (e.g., SOS or no SOS). | |
| 42723 | 100 | |
| 101 | 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 | 102 | 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 | 103 | slicing is disabled (e.g., by the minimizer). *) | 
| 42710 
84fcce345b5d
further improved type system setup based on Judgment Days
 blanchet parents: 
42709diff
changeset | 104 | |
| 38061 
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
 blanchet parents: 
38049diff
changeset | 105 | val known_perl_failures = | 
| 38094 | 106 | [(CantConnect, "HTTP error"), | 
| 107 | (NoPerl, "env: perl"), | |
| 38065 | 108 | (NoLibwwwPerl, "Can't locate HTTP")] | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 109 | |
| 45203 | 110 | fun known_szs_failures wrap = | 
| 111 | [(Unprovable, wrap "CounterSatisfiable"), | |
| 112 | (Unprovable, wrap "Satisfiable"), | |
| 113 | (GaveUp, wrap "GaveUp"), | |
| 114 | (GaveUp, wrap "Unknown"), | |
| 115 | (GaveUp, wrap "Incomplete"), | |
| 116 | (ProofMissing, wrap "Theorem"), | |
| 117 | (ProofMissing, wrap "Unsatisfiable"), | |
| 118 | (TimedOut, wrap "Timeout"), | |
| 119 | (Inappropriate, wrap "Inappropriate"), | |
| 120 | (OutOfResources, wrap "ResourceOut"), | |
| 121 | (OutOfResources, wrap "MemoryOut"), | |
| 122 | (Interrupted, wrap "Forced"), | |
| 123 | (Interrupted, wrap "User")] | |
| 124 | ||
| 125 | val known_szs_status_failures = known_szs_failures (prefix "SZS status ") | |
| 126 | val known_says_failures = known_szs_failures (prefix " says ") | |
| 127 | ||
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 128 | (* named ATPs *) | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 129 | |
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 130 | val eN = "e" | 
| 44590 | 131 | val e_sineN = "e_sine" | 
| 132 | val e_tofofN = "e_tofof" | |
| 45338 | 133 | val iproverN = "iprover" | 
| 134 | val iprover_eqN = "iprover_eq" | |
| 44099 | 135 | val leo2N = "leo2" | 
| 45365 | 136 | val dummy_tff1N = "dummy_tff1" (* experimental *) | 
| 137 | val dummy_thfN = "dummy_thf" (* experimental *) | |
| 44099 | 138 | val satallaxN = "satallax" | 
| 44590 | 139 | val snarkN = "snark" | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 140 | val spassN = "spass" | 
| 45365 | 141 | val spass_newN = "spass_new" (* experimental *) | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 142 | val vampireN = "vampire" | 
| 44590 | 143 | val waldmeisterN = "waldmeister" | 
| 44423 
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
 blanchet parents: 
44422diff
changeset | 144 | val z3_tptpN = "z3_tptp" | 
| 40060 
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
 blanchet parents: 
40059diff
changeset | 145 | val remote_prefix = "remote_" | 
| 38001 
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
 blanchet parents: 
38000diff
changeset | 146 | |
| 38023 | 147 | structure Data = Theory_Data | 
| 148 | ( | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 149 | type T = (atp_config * stamp) Symtab.table | 
| 38023 | 150 | val empty = Symtab.empty | 
| 151 | val extend = I | |
| 152 | fun merge data : T = Symtab.merge (eq_snd op =) data | |
| 153 |     handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".")
 | |
| 154 | ) | |
| 38017 
3ad3e3ca2451
move Sledgehammer-specific code out of "Sledgehammer_TPTP_Format"
 blanchet parents: 
38015diff
changeset | 155 | |
| 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 | 156 | 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 | 157 | |
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 158 | val sosN = "sos" | 
| 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 159 | val no_sosN = "no_sos" | 
| 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 160 | |
| 44099 | 161 | val force_sos = Attrib.setup_config_bool @{binding atp_force_sos} (K false)
 | 
| 162 | ||
| 39491 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39375diff
changeset | 163 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 164 | (* E *) | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 165 | |
| 44420 | 166 | fun is_old_e_version () = (string_ord (getenv "E_VERSION", "1.2w") = LESS) | 
| 167 | ||
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 168 | val tstp_proof_delims = | 
| 42962 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 blanchet parents: 
42955diff
changeset | 169 |   [("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation"),
 | 
| 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 blanchet parents: 
42955diff
changeset | 170 |    ("% SZS output start CNFRefutation", "% SZS output end CNFRefutation")]
 | 
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 171 | |
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 172 | val e_smartN = "smart" | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 173 | val e_autoN = "auto" | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 174 | val e_fun_weightN = "fun_weight" | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 175 | val e_sym_offset_weightN = "sym_offset_weight" | 
| 41725 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 176 | |
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 177 | val e_weight_method = | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 178 |   Attrib.setup_config_string @{binding atp_e_weight_method} (K e_smartN)
 | 
| 41770 | 179 | (* FUDGE *) | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 180 | val e_default_fun_weight = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 181 |   Attrib.setup_config_real @{binding atp_e_default_fun_weight} (K 20.0)
 | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 182 | val e_fun_weight_base = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 183 |   Attrib.setup_config_real @{binding atp_e_fun_weight_base} (K 0.0)
 | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 184 | val e_fun_weight_span = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 185 |   Attrib.setup_config_real @{binding atp_e_fun_weight_span} (K 40.0)
 | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 186 | val e_default_sym_offs_weight = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 187 |   Attrib.setup_config_real @{binding atp_e_default_sym_offs_weight} (K 1.0)
 | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 188 | val e_sym_offs_weight_base = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 189 |   Attrib.setup_config_real @{binding atp_e_sym_offs_weight_base} (K ~20.0)
 | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 190 | val e_sym_offs_weight_span = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 191 |   Attrib.setup_config_real @{binding atp_e_sym_offs_weight_span} (K 60.0)
 | 
| 41725 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 192 | |
| 42443 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42332diff
changeset | 193 | fun e_weight_method_case method fw sow = | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 194 | if method = e_fun_weightN then fw | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 195 | else if method = e_sym_offset_weightN then sow | 
| 43478 | 196 |   else raise Fail ("unexpected " ^ quote method)
 | 
| 41725 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 197 | |
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 198 | fun scaled_e_weight ctxt method w = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 199 | w * Config.get ctxt | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 200 | (e_weight_method_case method e_fun_weight_span e_sym_offs_weight_span) | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 201 | + Config.get ctxt | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 202 | (e_weight_method_case method e_fun_weight_base e_sym_offs_weight_base) | 
| 41725 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 203 | |> 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 | 204 | |
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 205 | fun e_weight_arguments ctxt method weights = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 206 | if method = e_autoN then | 
| 41725 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 207 | "-xAutoDev" | 
| 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 208 | else | 
| 43622 | 209 | (* supplied by Stephan Schulz *) | 
| 41314 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 210 | "--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 | 211 | \--destructive-er-aggressive --destructive-er --presat-simplify \ | 
| 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 212 | \--prefer-initial-clauses -tKBO6 -winvfreqrank -c1 -Ginvfreqconjmax -F1 \ | 
| 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 213 | \--delete-bad-limit=150000000 -WSelectMaxLComplexAvoidPosPred \ | 
| 42443 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42332diff
changeset | 214 | \-H'(4*" ^ e_weight_method_case method "FunWeight" "SymOffsetWeight" ^ | 
| 41725 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 215 | "(SimulateSOS, " ^ | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 216 | (e_weight_method_case method e_default_fun_weight e_default_sym_offs_weight | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 217 | |> Config.get ctxt |> Real.ceil |> signed_string_of_int) ^ | 
| 41314 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 218 | ",20,1.5,1.5,1" ^ | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 219 | (weights () | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 220 | |> map (fn (s, w) => "," ^ s ^ ":" ^ scaled_e_weight ctxt method w) | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 221 | |> implode) ^ | 
| 41314 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 222 | "),3*ConjectureGeneralSymbolWeight(PreferNonGoals,200,100,200,50,50,1,100,\ | 
| 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 223 | \1.5,1.5,1),1*Clauseweight(PreferProcessed,1,1,1),1*\ | 
| 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 224 | \FIFOWeight(PreferProcessed))'" | 
| 41313 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 blanchet parents: 
41269diff
changeset | 225 | |
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 226 | fun effective_e_weight_method ctxt = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 227 | if is_old_e_version () then e_autoN else Config.get ctxt e_weight_method | 
| 42443 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42332diff
changeset | 228 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 229 | val e_config : atp_config = | 
| 38092 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38090diff
changeset | 230 |   {exec = ("E_HOME", "eproof"),
 | 
| 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38090diff
changeset | 231 | required_execs = [], | 
| 43354 
396aaa15dd8b
pass --trim option to "eproof" script to speed up proof reconstruction
 blanchet parents: 
43288diff
changeset | 232 | arguments = | 
| 43567 
dda3e38cc351
remove experimental trimming feature -- it slowed down things on Linux for some reason
 blanchet parents: 
43566diff
changeset | 233 | fn ctxt => fn _ => fn method => fn timeout => fn weights => | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 234 | "--tstp-in --tstp-out -l5 " ^ e_weight_arguments ctxt method weights ^ | 
| 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 | 235 | " -tAutoDev --silent --cpu-limit=" ^ string_of_int (to_secs 2 timeout), | 
| 42962 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 blanchet parents: 
42955diff
changeset | 236 | proof_delims = tstp_proof_delims, | 
| 36265 
41c9e755e552
distinguish between the different ATP errors in the user interface;
 blanchet parents: 
36264diff
changeset | 237 | known_failures = | 
| 45203 | 238 | known_szs_status_failures @ | 
| 239 | [(TimedOut, "Failure: Resource limit exceeded (time)"), | |
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 240 | (TimedOut, "time limit exceeded"), | 
| 45203 | 241 | (OutOfResources, "# Cannot determine problem status")], | 
| 43466 
52f040bcfae7
tweaked TPTP formula kind for typing information used in the conjecture
 blanchet parents: 
43354diff
changeset | 242 | conj_sym_kind = Hypothesis, | 
| 42709 
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
 blanchet parents: 
42708diff
changeset | 243 | prem_kind = Conjecture, | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 244 | best_slices = fn ctxt => | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 245 | let val method = effective_e_weight_method ctxt in | 
| 43474 
423cd1ecf714
optimized E's time slicing, based on latest exhaustive Judgment Day results
 blanchet parents: 
43473diff
changeset | 246 | (* FUDGE *) | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 247 | if method = e_smartN then | 
| 45521 | 248 | [(0.333, (true, (500, FOF, "mono_tags??", combinatorsN, | 
| 249 | e_fun_weightN))), | |
| 250 | (0.334, (true, (50, FOF, "mono_guards??", combinatorsN, | |
| 251 | e_fun_weightN))), | |
| 252 | (0.333, (true, (1000, FOF, "mono_tags??", combinatorsN, | |
| 253 | e_sym_offset_weightN)))] | |
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 254 | else | 
| 45521 | 255 | [(1.0, (true, (500, FOF, "mono_tags??", combinatorsN, method)))] | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 256 | end} | 
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 257 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 258 | val e = (eN, e_config) | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 259 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 260 | |
| 44099 | 261 | (* LEO-II *) | 
| 262 | ||
| 44754 | 263 | val leo2_thf0 = THF (TPTP_Monomorphic, TPTP_Explicit, THF_Without_Choice) | 
| 264 | ||
| 44099 | 265 | val leo2_config : atp_config = | 
| 266 |   {exec = ("LEO2_HOME", "leo"),
 | |
| 267 | required_execs = [], | |
| 268 | arguments = | |
| 269 | fn _ => fn _ => fn sos => fn timeout => fn _ => | |
| 45300 
d8c8c2fcab2c
specify proof output level 1 (i.e. no detailed, potentially huge E proofs) to LEO-II; requires version 1.2.9
 blanchet parents: 
45234diff
changeset | 270 | "--proofoutput 1 --timeout " ^ string_of_int (to_secs 1 timeout) | 
| 44099 | 271 | |> sos = sosN ? prefix "--sos ", | 
| 272 | proof_delims = tstp_proof_delims, | |
| 45207 | 273 | known_failures = | 
| 274 | known_szs_status_failures @ | |
| 275 | [(TimedOut, "CPU time limit exceeded, terminating")], | |
| 44099 | 276 | conj_sym_kind = Axiom, | 
| 277 | prem_kind = Hypothesis, | |
| 278 | best_slices = fn ctxt => | |
| 279 | (* FUDGE *) | |
| 45521 | 280 | [(0.667, (false, (150, leo2_thf0, "mono_simple_higher", lam_liftingN, | 
| 281 | sosN))), | |
| 282 | (0.333, (true, (50, leo2_thf0, "mono_simple_higher", lam_liftingN, | |
| 283 | no_sosN)))] | |
| 44099 | 284 | |> (if Config.get ctxt force_sos then hd #> apfst (K 1.0) #> single | 
| 285 | else I)} | |
| 39491 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39375diff
changeset | 286 | |
| 44099 | 287 | val leo2 = (leo2N, leo2_config) | 
| 288 | ||
| 289 | ||
| 290 | (* Satallax *) | |
| 291 | ||
| 44754 | 292 | val satallax_thf0 = THF (TPTP_Monomorphic, TPTP_Explicit, THF_With_Choice) | 
| 293 | ||
| 44099 | 294 | val satallax_config : atp_config = | 
| 295 |   {exec = ("SATALLAX_HOME", "satallax"),
 | |
| 296 | required_execs = [], | |
| 297 | arguments = | |
| 298 | fn _ => fn _ => fn _ => fn timeout => fn _ => | |
| 45162 | 299 | "-p hocore -t " ^ string_of_int (to_secs 1 timeout), | 
| 300 | proof_delims = | |
| 301 |      [("% Higher-Order Unsat Core BEGIN", "% Higher-Order Unsat Core END")],
 | |
| 45203 | 302 | known_failures = known_szs_status_failures, | 
| 44099 | 303 | conj_sym_kind = Axiom, | 
| 304 | prem_kind = Hypothesis, | |
| 44416 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 blanchet parents: 
44391diff
changeset | 305 | best_slices = | 
| 44754 | 306 | (* FUDGE *) | 
| 45521 | 307 | K [(1.0, (true, (100, satallax_thf0, "mono_simple_higher", keep_lamsN, | 
| 308 | "")))]} | |
| 44099 | 309 | |
| 310 | val satallax = (satallaxN, satallax_config) | |
| 311 | ||
| 312 | ||
| 313 | (* 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 | 314 | |
| 36219 
16670b4f0baa
set SPASS option on the command-line, so that it doesn't vanish when moving to TPTP format
 blanchet parents: 
36190diff
changeset | 315 | (* The "-VarWeight=3" option helps the higher-order problems, probably by | 
| 44450 
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
 blanchet parents: 
44425diff
changeset | 316 | counteracting the presence of explicit application operators. *) | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 317 | val spass_config : atp_config = | 
| 38092 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38090diff
changeset | 318 |   {exec = ("ISABELLE_ATP", "scripts/spass"),
 | 
| 39002 | 319 |    required_execs = [("SPASS_HOME", "SPASS"), ("SPASS_HOME", "tptp2dfg")],
 | 
| 43569 
b342cd125533
removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
 blanchet parents: 
43567diff
changeset | 320 | arguments = fn _ => fn _ => fn sos => fn timeout => fn _ => | 
| 37962 
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
 blanchet parents: 
37926diff
changeset | 321 |      ("-Auto -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \
 | 
| 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 | 322 | \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_secs 1 timeout)) | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 323 | |> sos = sosN ? prefix "-SOS=1 ", | 
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 324 |    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 | 325 | known_failures = | 
| 38061 
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
 blanchet parents: 
38049diff
changeset | 326 | known_perl_failures @ | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
42999diff
changeset | 327 | [(GaveUp, "SPASS beiseite: Completion found"), | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 328 | (TimedOut, "SPASS beiseite: Ran out of time"), | 
| 36965 | 329 | (OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"), | 
| 37413 | 330 | (MalformedInput, "Undefined symbol"), | 
| 37414 
d0cea0796295
expect SPASS 3.7, and give a friendly warning if an older version is used
 blanchet parents: 
37413diff
changeset | 331 | (MalformedInput, "Free Variable"), | 
| 44391 | 332 | (Unprovable, "No formulae and clauses found in input file"), | 
| 39263 
e2a3c435334b
more precise error messages when Vampire is interrupted or SPASS runs into an internal bug
 blanchet parents: 
39262diff
changeset | 333 | (InternalError, "Please report this error")], | 
| 43466 
52f040bcfae7
tweaked TPTP formula kind for typing information used in the conjecture
 blanchet parents: 
43354diff
changeset | 334 | conj_sym_kind = Hypothesis, | 
| 42709 
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
 blanchet parents: 
42708diff
changeset | 335 | prem_kind = Conjecture, | 
| 42725 
64dea91bbe0e
added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
 blanchet parents: 
42723diff
changeset | 336 | best_slices = fn ctxt => | 
| 42723 | 337 | (* FUDGE *) | 
| 45521 | 338 | [(0.333, (false, (150, DFG DFG_Unsorted, "mono_tags??", lam_liftingN, | 
| 339 | sosN))), | |
| 340 | (0.333, (false, (300, DFG DFG_Unsorted, "poly_tags??", lam_liftingN, | |
| 341 | sosN))), | |
| 45876 
40952db4e57b
SPASS is incomplete because of the -Splits and -FullRed options, not just because of -SOS=1 -- don't pretend the opposite
 blanchet parents: 
45521diff
changeset | 342 | (0.334, (false, (50, DFG DFG_Unsorted, "mono_tags??", lam_liftingN, | 
| 
40952db4e57b
SPASS is incomplete because of the -Splits and -FullRed options, not just because of -SOS=1 -- don't pretend the opposite
 blanchet parents: 
45521diff
changeset | 343 | no_sosN)))] | 
| 44099 | 344 | |> (if Config.get ctxt force_sos then hd #> apfst (K 1.0) #> single | 
| 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 | 345 | else I)} | 
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 346 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 347 | val spass = (spassN, spass_config) | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 348 | |
| 45301 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 349 | (* Experimental *) | 
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 350 | val spass_new_config : atp_config = | 
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 351 |   {exec = ("SPASS_HOME", "SPASS"),
 | 
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 352 | required_execs = [], | 
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 353 | arguments = #arguments spass_config, | 
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 354 | proof_delims = #proof_delims spass_config, | 
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 355 | known_failures = #known_failures spass_config, | 
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 356 | conj_sym_kind = #conj_sym_kind spass_config, | 
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 357 | prem_kind = #prem_kind spass_config, | 
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 358 | best_slices = fn ctxt => | 
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 359 | (* FUDGE *) | 
| 45521 | 360 | [(0.333, (false, (150, DFG DFG_Sorted, "mono_simple", lam_liftingN, | 
| 361 | sosN))) (*, | |
| 362 | (0.333, (false, (300, DFG DFG_Sorted, "poly_tags??", lam_liftingN, | |
| 363 | sosN))), | |
| 364 | (0.334, (true, (50, DFG DFG_Sorted, "mono_simple", lam_liftingN, | |
| 365 | no_sosN)))*)] | |
| 45301 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 366 | |> (if Config.get ctxt force_sos then hd #> apfst (K 1.0) #> single | 
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 367 | else I)} | 
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 368 | |
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 369 | val spass_new = (spass_newN, spass_new_config) | 
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 370 | |
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 371 | |
| 37509 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 372 | (* Vampire *) | 
| 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 373 | |
| 44507 
e08158671ef4
disable TFF for Vampire 1.8 until they've fixed the soundness issues and it's back on SystemOnTPTP
 blanchet parents: 
44503diff
changeset | 374 | (* Vampire 1.8 has TFF support, but it's buggy and therefore disabled on | 
| 
e08158671ef4
disable TFF for Vampire 1.8 until they've fixed the soundness issues and it's back on SystemOnTPTP
 blanchet parents: 
44503diff
changeset | 375 | SystemOnTPTP. *) | 
| 44420 | 376 | fun is_old_vampire_version () = | 
| 44507 
e08158671ef4
disable TFF for Vampire 1.8 until they've fixed the soundness issues and it's back on SystemOnTPTP
 blanchet parents: 
44503diff
changeset | 377 | string_ord (getenv "VAMPIRE_VERSION", "1.8") <> GREATER | 
| 44420 | 378 | |
| 44754 | 379 | val vampire_tff0 = TFF (TPTP_Monomorphic, TPTP_Implicit) | 
| 44589 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 blanchet parents: 
44586diff
changeset | 380 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 381 | val vampire_config : atp_config = | 
| 38092 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38090diff
changeset | 382 |   {exec = ("VAMPIRE_HOME", "vampire"),
 | 
| 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38090diff
changeset | 383 | required_execs = [], | 
| 43569 
b342cd125533
removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
 blanchet parents: 
43567diff
changeset | 384 | arguments = fn _ => fn _ => fn sos => fn timeout => fn _ => | 
| 44417 | 385 | "--mode casc -t " ^ string_of_int (to_secs 1 timeout) ^ | 
| 45234 
5509362b924b
disable Vampire's BDD optimization, which sometimes yields so huge proofs that this causes problems for reconstruction
 blanchet parents: 
45207diff
changeset | 386 | " --proof tptp --output_axiom_names on\ | 
| 
5509362b924b
disable Vampire's BDD optimization, which sometimes yields so huge proofs that this causes problems for reconstruction
 blanchet parents: 
45207diff
changeset | 387 | \ --forced_options propositional_to_bdd=off\ | 
| 44417 | 388 | \ --thanks \"Andrei and Krystof\" --input_file" | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 389 | |> 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 | 390 | proof_delims = | 
| 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 391 |      [("=========== Refutation ==========",
 | 
| 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 392 | "======= End of refutation ======="), | 
| 38033 | 393 |       ("% SZS output start Refutation", "% SZS output end Refutation"),
 | 
| 394 |       ("% SZS output start Proof", "% SZS output end Proof")],
 | |
| 37509 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 395 | known_failures = | 
| 45203 | 396 | known_szs_status_failures @ | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
42999diff
changeset | 397 | [(GaveUp, "UNPROVABLE"), | 
| 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
42999diff
changeset | 398 | (GaveUp, "CANNOT PROVE"), | 
| 37509 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 399 | (Unprovable, "Satisfiability detected"), | 
| 38647 
5500241da479
play with fudge factor + parse one more Vampire error
 blanchet parents: 
38646diff
changeset | 400 | (Unprovable, "Termination reason: Satisfiable"), | 
| 39263 
e2a3c435334b
more precise error messages when Vampire is interrupted or SPASS runs into an internal bug
 blanchet parents: 
39262diff
changeset | 401 | (Interrupted, "Aborted by signal SIGINT")], | 
| 43466 
52f040bcfae7
tweaked TPTP formula kind for typing information used in the conjecture
 blanchet parents: 
43354diff
changeset | 402 | conj_sym_kind = Conjecture, | 
| 42709 
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
 blanchet parents: 
42708diff
changeset | 403 | prem_kind = Conjecture, | 
| 42725 
64dea91bbe0e
added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
 blanchet parents: 
42723diff
changeset | 404 | best_slices = fn ctxt => | 
| 42723 | 405 | (* FUDGE *) | 
| 44420 | 406 | (if is_old_vampire_version () then | 
| 45521 | 407 | [(0.333, (false, (150, FOF, "poly_guards??", hybrid_lamsN, sosN))), | 
| 408 | (0.333, (false, (500, FOF, "mono_tags??", hybrid_lamsN, sosN))), | |
| 409 | (0.334, (true, (50, FOF, "mono_guards??", hybrid_lamsN, no_sosN)))] | |
| 44420 | 410 | else | 
| 45521 | 411 | [(0.333, (false, (150, vampire_tff0, "poly_guards??", hybrid_lamsN, | 
| 412 | sosN))), | |
| 413 | (0.333, (false, (500, vampire_tff0, "mono_simple", hybrid_lamsN, | |
| 414 | sosN))), | |
| 415 | (0.334, (true, (50, vampire_tff0, "mono_simple", hybrid_lamsN, | |
| 416 | no_sosN)))]) | |
| 44099 | 417 | |> (if Config.get ctxt force_sos then hd #> apfst (K 1.0) #> single | 
| 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 | 418 | else I)} | 
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 419 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 420 | val vampire = (vampireN, vampire_config) | 
| 37509 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 421 | |
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 422 | |
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 423 | (* Z3 with TPTP syntax *) | 
| 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 424 | |
| 44754 | 425 | val z3_tff0 = TFF (TPTP_Monomorphic, TPTP_Implicit) | 
| 44589 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 blanchet parents: 
44586diff
changeset | 426 | |
| 44423 
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
 blanchet parents: 
44422diff
changeset | 427 | val z3_tptp_config : atp_config = | 
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 428 |   {exec = ("Z3_HOME", "z3"),
 | 
| 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 429 | required_execs = [], | 
| 43354 
396aaa15dd8b
pass --trim option to "eproof" script to speed up proof reconstruction
 blanchet parents: 
43288diff
changeset | 430 | arguments = fn _ => fn _ => fn _ => fn timeout => fn _ => | 
| 44420 | 431 | "MBQI=true -tptp -t:" ^ string_of_int (to_secs 1 timeout), | 
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 432 | proof_delims = [], | 
| 45203 | 433 | known_failures = known_szs_status_failures, | 
| 42709 
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
 blanchet parents: 
42708diff
changeset | 434 | conj_sym_kind = Hypothesis, | 
| 
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
 blanchet parents: 
42708diff
changeset | 435 | prem_kind = Hypothesis, | 
| 42723 | 436 | best_slices = | 
| 44423 
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
 blanchet parents: 
44422diff
changeset | 437 | (* FUDGE *) | 
| 45521 | 438 | K [(0.5, (false, (250, z3_tff0, "mono_simple", combinatorsN, ""))), | 
| 439 | (0.25, (false, (125, z3_tff0, "mono_simple", combinatorsN, ""))), | |
| 440 | (0.125, (false, (62, z3_tff0, "mono_simple", combinatorsN, ""))), | |
| 441 | (0.125, (false, (31, z3_tff0, "mono_simple", combinatorsN, "")))]} | |
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 442 | |
| 44423 
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
 blanchet parents: 
44422diff
changeset | 443 | val z3_tptp = (z3_tptpN, z3_tptp_config) | 
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 444 | |
| 44590 | 445 | |
| 44754 | 446 | (* Not really a prover: Experimental Polymorphic TFF and THF output *) | 
| 44590 | 447 | |
| 44754 | 448 | fun dummy_config format type_enc : atp_config = | 
| 44596 | 449 |   {exec = ("ISABELLE_ATP", "scripts/dummy_atp"),
 | 
| 44590 | 450 | required_execs = [], | 
| 451 | arguments = K (K (K (K (K "")))), | |
| 452 | proof_delims = [], | |
| 45203 | 453 | known_failures = known_szs_status_failures, | 
| 44590 | 454 | conj_sym_kind = Hypothesis, | 
| 455 | prem_kind = Hypothesis, | |
| 45521 | 456 | best_slices = | 
| 457 | K [(1.0, (false, (200, format, type_enc, | |
| 458 | if is_format_higher_order format then keep_lamsN | |
| 459 | else combinatorsN, "")))]} | |
| 44590 | 460 | |
| 45365 | 461 | val dummy_tff1_format = TFF (TPTP_Polymorphic, TPTP_Explicit) | 
| 462 | val dummy_tff1_config = dummy_config dummy_tff1_format "poly_simple" | |
| 463 | val dummy_tff1 = (dummy_tff1N, dummy_tff1_config) | |
| 44590 | 464 | |
| 45365 | 465 | val dummy_thf_format = THF (TPTP_Polymorphic, TPTP_Explicit, THF_With_Choice) | 
| 466 | val dummy_thf_config = dummy_config dummy_thf_format "poly_simple_higher" | |
| 467 | val dummy_thf = (dummy_thfN, dummy_thf_config) | |
| 44754 | 468 | |
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 469 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 470 | (* Remote ATP invocation via SystemOnTPTP *) | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 471 | |
| 38061 
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
 blanchet parents: 
38049diff
changeset | 472 | val systems = Synchronized.var "atp_systems" ([] : string list) | 
| 31835 | 473 | |
| 474 | fun get_systems () = | |
| 44589 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 blanchet parents: 
44586diff
changeset | 475 | case Isabelle_System.bash_output | 
| 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 blanchet parents: 
44586diff
changeset | 476 | "\"$ISABELLE_ATP/scripts/remote_atp\" -w 2>&1" of | 
| 39491 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39375diff
changeset | 477 | (output, 0) => split_lines output | 
| 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39375diff
changeset | 478 | | (output, _) => | 
| 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39375diff
changeset | 479 | error (case extract_known_failure known_perl_failures output of | 
| 41744 | 480 | SOME failure => string_for_failure failure | 
| 39491 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39375diff
changeset | 481 | | NONE => perhaps (try (unsuffix "\n")) output ^ ".") | 
| 31835 | 482 | |
| 42537 
25ceb855a18b
improve version handling -- prefer versions of ToFoF, SInE, and SNARK that are known to work
 blanchet parents: 
42535diff
changeset | 483 | fun find_system name [] systems = | 
| 
25ceb855a18b
improve version handling -- prefer versions of ToFoF, SInE, and SNARK that are known to work
 blanchet parents: 
42535diff
changeset | 484 | find_first (String.isPrefix (name ^ "---")) 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 | 485 | | find_system name (version :: versions) systems = | 
| 
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
 blanchet parents: 
38685diff
changeset | 486 | case find_first (String.isPrefix (name ^ "---" ^ version)) systems of | 
| 
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
 blanchet parents: 
38685diff
changeset | 487 | NONE => find_system name versions systems | 
| 
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
 blanchet parents: 
38685diff
changeset | 488 | | 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 | 489 | |
| 
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
 blanchet parents: 
38685diff
changeset | 490 | fun get_system name versions = | 
| 38589 
b03f8fe043ec
added "max_relevant_per_iter" option to Sledgehammer
 blanchet parents: 
38588diff
changeset | 491 | Synchronized.change_result systems | 
| 
b03f8fe043ec
added "max_relevant_per_iter" option to Sledgehammer
 blanchet parents: 
38588diff
changeset | 492 | (fn systems => (if null systems then get_systems () else systems) | 
| 42955 | 493 | |> `(`(find_system name versions))) | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 494 | |
| 38690 
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
 blanchet parents: 
38685diff
changeset | 495 | fun the_system name versions = | 
| 
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
 blanchet parents: 
38685diff
changeset | 496 | case get_system name versions of | 
| 42955 | 497 | (SOME sys, _) => sys | 
| 498 |   | (NONE, []) => error ("SystemOnTPTP is currently not available.")
 | |
| 499 | | (NONE, syss) => | |
| 500 |     error ("System " ^ quote name ^ " is not available at SystemOnTPTP.\n" ^
 | |
| 501 | "(Available systems: " ^ commas_quote syss ^ ".)") | |
| 31835 | 502 | |
| 41148 | 503 | val max_remote_secs = 240 (* give Geoff Sutcliffe's servers a break *) | 
| 504 | ||
| 38690 
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
 blanchet parents: 
38685diff
changeset | 505 | fun remote_config system_name system_versions proof_delims known_failures | 
| 44416 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 blanchet parents: 
44391diff
changeset | 506 | conj_sym_kind prem_kind best_slice : atp_config = | 
| 38092 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38090diff
changeset | 507 |   {exec = ("ISABELLE_ATP", "scripts/remote_atp"),
 | 
| 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38090diff
changeset | 508 | required_execs = [], | 
| 43354 
396aaa15dd8b
pass --trim option to "eproof" script to speed up proof reconstruction
 blanchet parents: 
43288diff
changeset | 509 | arguments = fn _ => fn _ => fn _ => fn timeout => fn _ => | 
| 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 | 510 | "-t " ^ string_of_int (Int.min (max_remote_secs, to_secs 1 timeout)) | 
| 41148 | 511 | ^ " -s " ^ the_system system_name system_versions, | 
| 42962 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 blanchet parents: 
42955diff
changeset | 512 | proof_delims = union (op =) tstp_proof_delims proof_delims, | 
| 45203 | 513 | known_failures = known_failures @ known_perl_failures @ known_says_failures, | 
| 42709 
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
 blanchet parents: 
42708diff
changeset | 514 | conj_sym_kind = conj_sym_kind, | 
| 
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
 blanchet parents: 
42708diff
changeset | 515 | prem_kind = prem_kind, | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 516 | best_slices = fn ctxt => | 
| 45521 | 517 | let val (max_relevant, format, type_enc, lam_trans) = best_slice ctxt in | 
| 518 | [(1.0, (false, (max_relevant, format, type_enc, lam_trans, "")))] | |
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 519 | end} | 
| 42443 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42332diff
changeset | 520 | |
| 43500 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 521 | fun remotify_config system_name system_versions best_slice | 
| 44416 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 blanchet parents: 
44391diff
changeset | 522 |         ({proof_delims, known_failures, conj_sym_kind, prem_kind, ...}
 | 
| 43500 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 523 | : atp_config) : atp_config = | 
| 38690 
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
 blanchet parents: 
38685diff
changeset | 524 | remote_config system_name system_versions proof_delims known_failures | 
| 44416 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 blanchet parents: 
44391diff
changeset | 525 | conj_sym_kind prem_kind best_slice | 
| 38023 | 526 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 527 | fun remote_atp name system_name system_versions proof_delims known_failures | 
| 44416 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 blanchet parents: 
44391diff
changeset | 528 | conj_sym_kind prem_kind best_slice = | 
| 40060 
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
 blanchet parents: 
40059diff
changeset | 529 | (remote_prefix ^ name, | 
| 38690 
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
 blanchet parents: 
38685diff
changeset | 530 | remote_config system_name system_versions proof_delims known_failures | 
| 44416 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 blanchet parents: 
44391diff
changeset | 531 | conj_sym_kind prem_kind best_slice) | 
| 43500 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 532 | fun remotify_atp (name, config) system_name system_versions best_slice = | 
| 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 533 | (remote_prefix ^ name, | 
| 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 534 | remotify_config system_name system_versions best_slice config) | 
| 28592 | 535 | |
| 44754 | 536 | val explicit_tff0 = TFF (TPTP_Monomorphic, TPTP_Explicit) | 
| 44589 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 blanchet parents: 
44586diff
changeset | 537 | |
| 43500 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 538 | val remote_e = | 
| 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 539 | remotify_atp e "EP" ["1.0", "1.1", "1.2"] | 
| 45521 | 540 | (K (750, FOF, "mono_tags??", combinatorsN) (* FUDGE *)) | 
| 44099 | 541 | val remote_leo2 = | 
| 542 | remotify_atp leo2 "LEO-II" ["1.2.8", "1.2.6"] | |
| 45521 | 543 | (K (100, leo2_thf0, "mono_simple_higher", lam_liftingN) (* FUDGE *)) | 
| 44099 | 544 | val remote_satallax = | 
| 545 | remotify_atp satallax "Satallax" ["2.1", "2.0", "2"] | |
| 45521 | 546 | (K (100, satallax_thf0, "mono_simple_higher", keep_lamsN) (* FUDGE *)) | 
| 43500 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 547 | val remote_vampire = | 
| 44499 | 548 | remotify_atp vampire "Vampire" ["1.8"] | 
| 45521 | 549 | (K (250, FOF, "mono_guards??", hybrid_lamsN) (* FUDGE *)) | 
| 44423 
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
 blanchet parents: 
44422diff
changeset | 550 | val remote_z3_tptp = | 
| 44754 | 551 | remotify_atp z3_tptp "Z3" ["3.0"] | 
| 45521 | 552 | (K (250, z3_tff0, "mono_simple", combinatorsN) (* FUDGE *)) | 
| 44092 
bf489e54d7f8
renamed E wrappers for consistency with CASC conventions
 blanchet parents: 
43989diff
changeset | 553 | val remote_e_sine = | 
| 
bf489e54d7f8
renamed E wrappers for consistency with CASC conventions
 blanchet parents: 
43989diff
changeset | 554 | remote_atp e_sineN "SInE" ["0.4"] [] (#known_failures e_config) Axiom | 
| 45521 | 555 | Conjecture (K (500, FOF, "mono_guards??", combinatorsN) (* FUDGE *)) | 
| 45338 | 556 | val remote_iprover = | 
| 557 | remote_atp iproverN "iProver" [] [] [] Axiom Conjecture | |
| 45521 | 558 | (K (150, FOF, "mono_guards??", lam_liftingN) (* FUDGE *)) | 
| 45338 | 559 | val remote_iprover_eq = | 
| 560 | remote_atp iprover_eqN "iProver-Eq" [] [] [] Axiom Conjecture | |
| 45521 | 561 | (K (150, FOF, "mono_guards??", lam_liftingN) (* FUDGE *)) | 
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 562 | val remote_snark = | 
| 42939 | 563 | remote_atp snarkN "SNARK" ["20080805r029", "20080805r024"] | 
| 45521 | 564 |       [("refutation.", "end_refutation.")] [] Hypothesis Hypothesis
 | 
| 565 | (K (100, explicit_tff0, "mono_simple", lam_liftingN) (* FUDGE *)) | |
| 44092 
bf489e54d7f8
renamed E wrappers for consistency with CASC conventions
 blanchet parents: 
43989diff
changeset | 566 | val remote_e_tofof = | 
| 44589 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 blanchet parents: 
44586diff
changeset | 567 | remote_atp e_tofofN "ToFoF" ["0.1"] [] (#known_failures e_config) Axiom | 
| 45521 | 568 | Hypothesis | 
| 569 | (K (150, explicit_tff0, "mono_simple", lam_liftingN) (* FUDGE *)) | |
| 42938 | 570 | val remote_waldmeister = | 
| 571 | remote_atp waldmeisterN "Waldmeister" ["710"] | |
| 45521 | 572 |       [("#START OF PROOF", "Proved Goals:")]
 | 
| 573 | [(OutOfResources, "Too many function symbols"), | |
| 574 | (Crashed, "Unrecoverable Segmentation Fault")] | |
| 575 | Hypothesis Hypothesis | |
| 576 | (K (50, CNF_UEQ, "mono_tags??", combinatorsN) (* FUDGE *)) | |
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 577 | |
| 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 578 | (* Setup *) | 
| 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 579 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 580 | fun add_atp (name, config) thy = | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 581 | Data.map (Symtab.update_new (name, (config, stamp ()))) thy | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 582 |   handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".")
 | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 583 | |
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 584 | fun get_atp thy name = | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 585 | the (Symtab.lookup (Data.get thy) name) |> fst | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 586 |   handle Option.Option => error ("Unknown ATP: " ^ name ^ ".")
 | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 587 | |
| 41727 
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
 blanchet parents: 
41725diff
changeset | 588 | val supported_atps = Symtab.keys o Data.get | 
| 36371 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 589 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 590 | fun is_atp_installed thy name = | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 591 |   let val {exec, required_execs, ...} = get_atp thy name in
 | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 592 | forall (curry (op <>) "" o getenv o fst) (exec :: required_execs) | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 593 | end | 
| 36371 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 594 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 595 | fun refresh_systems_on_tptp () = | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 596 | Synchronized.change systems (fn _ => get_systems ()) | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 597 | |
| 42962 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 blanchet parents: 
42955diff
changeset | 598 | val atps = | 
| 45365 | 599 | [e, leo2, dummy_tff1, dummy_thf, satallax, spass, spass_new, vampire, z3_tptp, | 
| 600 | remote_e, remote_e_sine, remote_e_tofof, remote_iprover, remote_iprover_eq, | |
| 45339 | 601 | remote_leo2, remote_satallax, remote_vampire, remote_z3_tptp, remote_snark, | 
| 602 | remote_waldmeister] | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 603 | val setup = fold add_atp atps | 
| 35867 | 604 | |
| 28592 | 605 | end; |