| author | wenzelm | 
| Tue, 03 Apr 2012 16:10:34 +0200 | |
| changeset 47291 | 6a641856a0e9 | 
| parent 47149 | 97f8c6c88134 | 
| child 47499 | 4b0daca2bf88 | 
| 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 | 
| 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 | 12 | 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 | 13 | type failure = ATP_Proof.failure | 
| 38023 | 14 | |
| 46409 
d4754183ccce
made option available to users (mostly for experiments)
 blanchet parents: 
46407diff
changeset | 15 | type slice_spec = int * 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 = | 
| 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 | 17 |     {exec : string list * string,
 | 
| 
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 | 18 | required_vars : string list list, | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 19 | arguments : | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 20 | Proof.context -> bool -> string -> Time.time | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 21 | -> term_order * (unit -> (string * int) list) | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 22 | * (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 | 23 | proof_delims : (string * string) list, | 
| 
1eaf4d437d4c
define type system in ATP module so that ATPs can suggest a type system
 blanchet parents: 
42577diff
changeset | 24 | 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 | 25 | 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 | 26 | prem_kind : formula_kind, | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 27 | best_slices : | 
| 46409 
d4754183ccce
made option available to users (mostly for experiments)
 blanchet parents: 
46407diff
changeset | 28 | Proof.context -> (real * (bool * (slice_spec * string))) list} | 
| 38023 | 29 | |
| 44099 | 30 | val force_sos : bool Config.T | 
| 47032 | 31 | 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 | 32 | 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 | 33 | 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 | 34 | 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 | 35 | val e_sym_offset_weightN : string | 
| 47032 | 36 | val e_selection_heuristic : string Config.T | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 37 | val e_default_fun_weight : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 38 | val e_fun_weight_base : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 39 | val e_fun_weight_span : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 40 | val e_default_sym_offs_weight : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 41 | val e_sym_offs_weight_base : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 42 | val e_sym_offs_weight_span : real Config.T | 
| 46643 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 43 | val alt_ergoN : string | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 44 | val dummy_thfN : string | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 45 | val eN : string | 
| 44590 | 46 | val e_sineN : string | 
| 47 | val e_tofofN : string | |
| 45338 | 48 | val iproverN : string | 
| 49 | val iprover_eqN : string | |
| 44590 | 50 | val leo2N : string | 
| 51 | val satallaxN : string | |
| 52 | val snarkN : string | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 53 | val spassN : string | 
| 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 | 54 | val spass_oldN : string | 
| 45301 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 55 | val spass_newN : string | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 56 | val vampireN : string | 
| 42938 | 57 | val waldmeisterN : string | 
| 44423 
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
 blanchet parents: 
44422diff
changeset | 58 | val z3_tptpN : string | 
| 40060 
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
 blanchet parents: 
40059diff
changeset | 59 | 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 | 60 | 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 | 61 | string -> string -> string list -> (string * string) list | 
| 44416 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 blanchet parents: 
44391diff
changeset | 62 | -> (failure * string) list -> formula_kind -> formula_kind | 
| 47074 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 63 | -> (Proof.context -> slice_spec * string) -> string * atp_config | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 64 | val add_atp : string * atp_config -> theory -> theory | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 65 | val get_atp : theory -> string -> atp_config | 
| 41727 
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
 blanchet parents: 
41725diff
changeset | 66 | val supported_atps : theory -> string list | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 67 | val is_atp_installed : theory -> string -> bool | 
| 35867 | 68 | 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 | 69 | val effective_term_order : Proof.context -> string -> term_order | 
| 35867 | 70 | val setup : theory -> theory | 
| 28592 | 71 | end; | 
| 72 | ||
| 36376 | 73 | structure ATP_Systems : ATP_SYSTEMS = | 
| 28592 | 74 | struct | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 75 | |
| 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 | 76 | open ATP_Problem | 
| 39491 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39375diff
changeset | 77 | open ATP_Proof | 
| 46320 | 78 | open ATP_Problem_Generate | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 79 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 80 | (* ATP configuration *) | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 81 | |
| 46409 
d4754183ccce
made option available to users (mostly for experiments)
 blanchet parents: 
46407diff
changeset | 82 | type slice_spec = int * atp_format * string * string * bool | 
| 
d4754183ccce
made option available to users (mostly for experiments)
 blanchet parents: 
46407diff
changeset | 83 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 84 | type atp_config = | 
| 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 | 85 |   {exec : string list * string,
 | 
| 
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 | 86 | required_vars : string list list, | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 87 | arguments : | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 88 | Proof.context -> bool -> string -> Time.time | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 89 | -> term_order * (unit -> (string * int) list) | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 90 | * (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 | 91 | proof_delims : (string * string) list, | 
| 
1eaf4d437d4c
define type system in ATP module so that ATPs can suggest a type system
 blanchet parents: 
42577diff
changeset | 92 | 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 | 93 | 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 | 94 | prem_kind : formula_kind, | 
| 46409 
d4754183ccce
made option available to users (mostly for experiments)
 blanchet parents: 
46407diff
changeset | 95 | best_slices : Proof.context -> (real * (bool * (slice_spec * string))) list} | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 96 | |
| 42723 | 97 | (* "best_slices" must be found empirically, taking a wholistic approach since | 
| 46407 
30e9720cc0b9
optimization: slice caching in case two consecutive slices are nearly identical
 blanchet parents: 
46402diff
changeset | 98 | the ATPs are run in parallel. The "real" component gives the faction of the | 
| 46409 
d4754183ccce
made option available to users (mostly for experiments)
 blanchet parents: 
46407diff
changeset | 99 | time available given to the slice and should add up to 1.0. The first "bool" | 
| 42723 | 100 | 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 | 101 | preferred number of facts to pass; the first "string", the preferred type | 
| 45521 | 102 | system (which should be sound or quasi-sound); the second "string", the | 
| 46409 
d4754183ccce
made option available to users (mostly for experiments)
 blanchet parents: 
46407diff
changeset | 103 | preferred lambda translation scheme; the second "bool", whether uncurried | 
| 
d4754183ccce
made option available to users (mostly for experiments)
 blanchet parents: 
46407diff
changeset | 104 | aliased should be generated; the third "string", extra information to | 
| 45521 | 105 | the prover (e.g., SOS or no SOS). | 
| 42723 | 106 | |
| 107 | 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 | 108 | 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 | 109 | slicing is disabled (e.g., by the minimizer). *) | 
| 42710 
84fcce345b5d
further improved type system setup based on Judgment Days
 blanchet parents: 
42709diff
changeset | 110 | |
| 38061 
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
 blanchet parents: 
38049diff
changeset | 111 | val known_perl_failures = | 
| 38094 | 112 | [(CantConnect, "HTTP error"), | 
| 113 | (NoPerl, "env: perl"), | |
| 38065 | 114 | (NoLibwwwPerl, "Can't locate HTTP")] | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 115 | |
| 45203 | 116 | fun known_szs_failures wrap = | 
| 117 | [(Unprovable, wrap "CounterSatisfiable"), | |
| 118 | (Unprovable, wrap "Satisfiable"), | |
| 119 | (GaveUp, wrap "GaveUp"), | |
| 120 | (GaveUp, wrap "Unknown"), | |
| 121 | (GaveUp, wrap "Incomplete"), | |
| 122 | (ProofMissing, wrap "Theorem"), | |
| 123 | (ProofMissing, wrap "Unsatisfiable"), | |
| 124 | (TimedOut, wrap "Timeout"), | |
| 125 | (Inappropriate, wrap "Inappropriate"), | |
| 126 | (OutOfResources, wrap "ResourceOut"), | |
| 127 | (OutOfResources, wrap "MemoryOut"), | |
| 128 | (Interrupted, wrap "Forced"), | |
| 129 | (Interrupted, wrap "User")] | |
| 130 | ||
| 131 | val known_szs_status_failures = known_szs_failures (prefix "SZS status ") | |
| 132 | val known_says_failures = known_szs_failures (prefix " says ") | |
| 133 | ||
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 134 | (* named ATPs *) | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 135 | |
| 46643 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 136 | val alt_ergoN = "alt_ergo" | 
| 47055 
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
 blanchet parents: 
47053diff
changeset | 137 | val dummy_thfN = "dummy_thf" (* for experiments *) | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 138 | val eN = "e" | 
| 44590 | 139 | val e_sineN = "e_sine" | 
| 140 | val e_tofofN = "e_tofof" | |
| 45338 | 141 | val iproverN = "iprover" | 
| 142 | val iprover_eqN = "iprover_eq" | |
| 44099 | 143 | val leo2N = "leo2" | 
| 144 | val satallaxN = "satallax" | |
| 44590 | 145 | val snarkN = "snark" | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 146 | val spassN = "spass" | 
| 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 | 147 | val spass_oldN = "spass_old" (* for experiments *) | 
| 
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 | 148 | val spass_newN = "spass_new" (* for experiments *) | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 149 | val vampireN = "vampire" | 
| 44590 | 150 | val waldmeisterN = "waldmeister" | 
| 44423 
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
 blanchet parents: 
44422diff
changeset | 151 | val z3_tptpN = "z3_tptp" | 
| 40060 
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
 blanchet parents: 
40059diff
changeset | 152 | val remote_prefix = "remote_" | 
| 38001 
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
 blanchet parents: 
38000diff
changeset | 153 | |
| 38023 | 154 | structure Data = Theory_Data | 
| 155 | ( | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 156 | type T = (atp_config * stamp) Symtab.table | 
| 38023 | 157 | val empty = Symtab.empty | 
| 158 | val extend = I | |
| 46407 
30e9720cc0b9
optimization: slice caching in case two consecutive slices are nearly identical
 blanchet parents: 
46402diff
changeset | 159 | fun merge data : T = | 
| 
30e9720cc0b9
optimization: slice caching in case two consecutive slices are nearly identical
 blanchet parents: 
46402diff
changeset | 160 | Symtab.merge (eq_snd (op =)) data | 
| 38023 | 161 |     handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".")
 | 
| 162 | ) | |
| 38017 
3ad3e3ca2451
move Sledgehammer-specific code out of "Sledgehammer_TPTP_Format"
 blanchet parents: 
38015diff
changeset | 163 | |
| 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 | 164 | 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 | 165 | |
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 166 | val sosN = "sos" | 
| 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 167 | val no_sosN = "no_sos" | 
| 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 168 | |
| 44099 | 169 | val force_sos = Attrib.setup_config_bool @{binding atp_force_sos} (K false)
 | 
| 170 | ||
| 47032 | 171 | val smartN = "smart" | 
| 47073 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 blanchet parents: 
47055diff
changeset | 172 | (* val kboN = "kbo" *) | 
| 47032 | 173 | val lpoN = "lpo" | 
| 47034 
77da780ddd6b
implement term order attribute (for experiments)
 blanchet parents: 
47033diff
changeset | 174 | val xweightsN = "_weights" | 
| 
77da780ddd6b
implement term order attribute (for experiments)
 blanchet parents: 
47033diff
changeset | 175 | val xprecN = "_prec" | 
| 
77da780ddd6b
implement term order attribute (for experiments)
 blanchet parents: 
47033diff
changeset | 176 | val xsimpN = "_simp" (* SPASS-specific *) | 
| 47032 | 177 | |
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 178 | (* Possible values for "atp_term_order": | 
| 47049 | 179 | "smart", "(kbo|lpo)(_weights)?(_prec|_simp)?" *) | 
| 47032 | 180 | val term_order = | 
| 181 |   Attrib.setup_config_string @{binding atp_term_order} (K smartN)
 | |
| 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_tff1 = TFF (TPTP_Polymorphic, TPTP_Explicit) | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 186 | |
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 187 | val alt_ergo_config : atp_config = | 
| 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 | 188 |   {exec = (["WHY3_HOME"], "why3"),
 | 
| 
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 | 189 | required_vars = [], | 
| 46643 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 190 | arguments = | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 191 | fn _ => fn _ => fn _ => fn timeout => fn _ => | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 192 | "--format tff1 --prover alt-ergo --timelimit " ^ | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 193 | string_of_int (to_secs 1 timeout), | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 194 | proof_delims = [], | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 195 | known_failures = | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 196 | [(ProofMissing, ": Valid"), | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 197 | (TimedOut, ": Timeout"), | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 198 | (GaveUp, ": Unknown")], | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 199 | conj_sym_kind = Hypothesis, | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 200 | prem_kind = Hypothesis, | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 201 | best_slices = fn _ => | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 202 | (* FUDGE *) | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 203 | [(1.0, (false, ((100, alt_ergo_tff1, "poly_native", liftingN, false), "")))]} | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 204 | |
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 205 | val alt_ergo = (alt_ergoN, alt_ergo_config) | 
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 206 | |
| 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 blanchet parents: 
46481diff
changeset | 207 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 208 | (* E *) | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 209 | |
| 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 | 210 | fun is_new_e_version () = (string_ord (getenv "E_VERSION", "1.2") = GREATER) | 
| 44420 | 211 | |
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 212 | val tstp_proof_delims = | 
| 42962 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 blanchet parents: 
42955diff
changeset | 213 |   [("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation"),
 | 
| 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 blanchet parents: 
42955diff
changeset | 214 |    ("% SZS output start CNFRefutation", "% SZS output end CNFRefutation")]
 | 
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 215 | |
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 216 | val e_smartN = "smart" | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 217 | val e_autoN = "auto" | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 218 | val e_fun_weightN = "fun_weight" | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 219 | val e_sym_offset_weightN = "sym_offset_weight" | 
| 41725 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 220 | |
| 47032 | 221 | val e_selection_heuristic = | 
| 222 |   Attrib.setup_config_string @{binding atp_e_selection_heuristic} (K e_smartN)
 | |
| 41770 | 223 | (* FUDGE *) | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 224 | val e_default_fun_weight = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 225 |   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 | 226 | val e_fun_weight_base = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 227 |   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 | 228 | val e_fun_weight_span = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 229 |   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 | 230 | val e_default_sym_offs_weight = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 231 |   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 | 232 | val e_sym_offs_weight_base = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 233 |   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 | 234 | val e_sym_offs_weight_span = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 235 |   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 | 236 | |
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 237 | fun e_selection_heuristic_case heuristic fw sow = | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 238 | if heuristic = e_fun_weightN then fw | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 239 | else if heuristic = e_sym_offset_weightN then sow | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 240 |   else raise Fail ("unexpected " ^ quote heuristic)
 | 
| 41725 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 241 | |
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 242 | fun scaled_e_selection_weight ctxt heuristic w = | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 243 | w * Config.get ctxt (e_selection_heuristic_case heuristic | 
| 47029 | 244 | e_fun_weight_span e_sym_offs_weight_span) | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 245 | + Config.get ctxt (e_selection_heuristic_case heuristic | 
| 47029 | 246 | e_fun_weight_base e_sym_offs_weight_base) | 
| 41725 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 247 | |> 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 | 248 | |
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 249 | fun e_selection_weight_arguments ctxt heuristic sel_weights = | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 250 | if heuristic = e_autoN then | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 251 | "-xAuto" | 
| 41725 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 252 | else | 
| 43622 | 253 | (* supplied by Stephan Schulz *) | 
| 41314 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 254 | "--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 | 255 | \--destructive-er-aggressive --destructive-er --presat-simplify \ | 
| 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 256 | \--prefer-initial-clauses -tKBO6 -winvfreqrank -c1 -Ginvfreqconjmax -F1 \ | 
| 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 257 | \--delete-bad-limit=150000000 -WSelectMaxLComplexAvoidPosPred \ | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 258 | \-H'(4*" ^ | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 259 | e_selection_heuristic_case heuristic "FunWeight" "SymOffsetWeight" ^ | 
| 41725 
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
 blanchet parents: 
41335diff
changeset | 260 | "(SimulateSOS, " ^ | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 261 | (e_selection_heuristic_case heuristic | 
| 47029 | 262 | e_default_fun_weight e_default_sym_offs_weight | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 263 | |> Config.get ctxt |> Real.ceil |> signed_string_of_int) ^ | 
| 41314 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 264 | ",20,1.5,1.5,1" ^ | 
| 47030 | 265 | (sel_weights () | 
| 47029 | 266 | |> map (fn (s, w) => "," ^ s ^ ":" ^ | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 267 | scaled_e_selection_weight ctxt heuristic w) | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 268 | |> implode) ^ | 
| 41314 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 269 | "),3*ConjectureGeneralSymbolWeight(PreferNonGoals,200,100,200,50,50,1,100,\ | 
| 
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
 blanchet parents: 
41313diff
changeset | 270 | \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 | 271 | \FIFOWeight(PreferProcessed))'" | 
| 41313 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 blanchet parents: 
41269diff
changeset | 272 | |
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 273 | val e_ord_weights = | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 274 | map (fn (s, w) => s ^ ":" ^ string_of_int w) #> space_implode "," | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 275 | fun e_ord_precedence [_] = "" | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 276 | | e_ord_precedence info = info |> map fst |> space_implode "<" | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 277 | |
| 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 | 278 | 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 | 279 | | e_term_order_info_arguments gen_weights gen_prec ord_info = | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 280 | let val ord_info = ord_info () in | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 281 | (if gen_weights then "--order-weights='" ^ e_ord_weights ord_info ^ "' " | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 282 | else "") ^ | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 283 | (if gen_prec then "--precedence='" ^ e_ord_precedence ord_info ^ "' " | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 284 | else "") | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 285 | end | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 286 | |
| 47032 | 287 | fun effective_e_selection_heuristic ctxt = | 
| 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 | 288 | if is_new_e_version () then Config.get ctxt e_selection_heuristic else e_autoN | 
| 42443 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42332diff
changeset | 289 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 290 | val e_config : atp_config = | 
| 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 | 291 |   {exec = (["E_HOME"], "eproof"),
 | 
| 
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 | 292 | required_vars = [], | 
| 43354 
396aaa15dd8b
pass --trim option to "eproof" script to speed up proof reconstruction
 blanchet parents: 
43288diff
changeset | 293 | arguments = | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 294 | fn ctxt => fn _ => fn heuristic => fn timeout => | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 295 |         fn ({is_lpo, gen_weights, gen_prec, ...}, ord_info, sel_weights) =>
 | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 296 | "--tstp-in --tstp-out --output-level=5 --silent " ^ | 
| 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 297 | e_selection_weight_arguments ctxt heuristic sel_weights ^ " " ^ | 
| 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 | 298 | e_term_order_info_arguments gen_weights gen_prec ord_info ^ " " ^ | 
| 47053 | 299 | "--term-ordering=" ^ (if is_lpo then "LPO4" else "KBO6") ^ " " ^ | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 300 | "--cpu-limit=" ^ string_of_int (to_secs 2 timeout), | 
| 42962 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 blanchet parents: 
42955diff
changeset | 301 | proof_delims = tstp_proof_delims, | 
| 36265 
41c9e755e552
distinguish between the different ATP errors in the user interface;
 blanchet parents: 
36264diff
changeset | 302 | known_failures = | 
| 45203 | 303 | known_szs_status_failures @ | 
| 304 | [(TimedOut, "Failure: Resource limit exceeded (time)"), | |
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 305 | (TimedOut, "time limit exceeded"), | 
| 45203 | 306 | (OutOfResources, "# Cannot determine problem status")], | 
| 43466 
52f040bcfae7
tweaked TPTP formula kind for typing information used in the conjecture
 blanchet parents: 
43354diff
changeset | 307 | 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 | 308 | prem_kind = Conjecture, | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42643diff
changeset | 309 | best_slices = fn ctxt => | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 310 | let val heuristic = effective_e_selection_heuristic ctxt in | 
| 43474 
423cd1ecf714
optimized E's time slicing, based on latest exhaustive Judgment Day results
 blanchet parents: 
43473diff
changeset | 311 | (* FUDGE *) | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 312 | if heuristic = e_smartN then | 
| 46449 | 313 | [(0.333, (true, ((500, FOF, "mono_tags??", combsN, false), e_fun_weightN))), | 
| 314 | (0.334, (true, ((50, FOF, "mono_guards??", combsN, false), e_fun_weightN))), | |
| 315 | (0.333, (true, ((1000, FOF, "mono_tags??", combsN, false), e_sym_offset_weightN)))] | |
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 316 | else | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
47034diff
changeset | 317 | [(1.0, (true, ((500, FOF, "mono_tags??", combsN, false), heuristic)))] | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 318 | end} | 
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 319 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 320 | val e = (eN, e_config) | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 321 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 322 | |
| 44099 | 323 | (* LEO-II *) | 
| 324 | ||
| 44754 | 325 | val leo2_thf0 = THF (TPTP_Monomorphic, TPTP_Explicit, THF_Without_Choice) | 
| 326 | ||
| 44099 | 327 | val leo2_config : atp_config = | 
| 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 | 328 |   {exec = (["LEO2_HOME"], "leo"),
 | 
| 
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 | 329 | required_vars = [], | 
| 44099 | 330 | arguments = | 
| 331 | 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 | 332 | "--proofoutput 1 --timeout " ^ string_of_int (to_secs 1 timeout) | 
| 44099 | 333 | |> sos = sosN ? prefix "--sos ", | 
| 334 | proof_delims = tstp_proof_delims, | |
| 45207 | 335 | known_failures = | 
| 336 | known_szs_status_failures @ | |
| 46481 | 337 | [(TimedOut, "CPU time limit exceeded, terminating"), | 
| 338 | (GaveUp, "No.of.Axioms")], | |
| 44099 | 339 | conj_sym_kind = Axiom, | 
| 340 | prem_kind = Hypothesis, | |
| 341 | best_slices = fn ctxt => | |
| 342 | (* FUDGE *) | |
| 46449 | 343 | [(0.667, (false, ((150, leo2_thf0, "mono_native_higher", liftingN, false), sosN))), | 
| 344 | (0.333, (true, ((50, leo2_thf0, "mono_native_higher", liftingN, false), no_sosN)))] | |
| 44099 | 345 | |> (if Config.get ctxt force_sos then hd #> apfst (K 1.0) #> single | 
| 346 | else I)} | |
| 39491 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39375diff
changeset | 347 | |
| 44099 | 348 | val leo2 = (leo2N, leo2_config) | 
| 349 | ||
| 350 | ||
| 351 | (* Satallax *) | |
| 352 | ||
| 44754 | 353 | val satallax_thf0 = THF (TPTP_Monomorphic, TPTP_Explicit, THF_With_Choice) | 
| 354 | ||
| 44099 | 355 | val satallax_config : atp_config = | 
| 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 | 356 |   {exec = (["SATALLAX_HOME"], "satallax"),
 | 
| 
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 | 357 | required_vars = [], | 
| 44099 | 358 | arguments = | 
| 359 | fn _ => fn _ => fn _ => fn timeout => fn _ => | |
| 45162 | 360 | "-p hocore -t " ^ string_of_int (to_secs 1 timeout), | 
| 361 | proof_delims = | |
| 362 |      [("% Higher-Order Unsat Core BEGIN", "% Higher-Order Unsat Core END")],
 | |
| 45203 | 363 | known_failures = known_szs_status_failures, | 
| 44099 | 364 | conj_sym_kind = Axiom, | 
| 365 | prem_kind = Hypothesis, | |
| 44416 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 blanchet parents: 
44391diff
changeset | 366 | best_slices = | 
| 44754 | 367 | (* FUDGE *) | 
| 46449 | 368 | K [(1.0, (true, ((100, satallax_thf0, "mono_native_higher", keep_lamsN, false), "")))]} | 
| 44099 | 369 | |
| 370 | val satallax = (satallaxN, satallax_config) | |
| 371 | ||
| 372 | ||
| 373 | (* 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 | 374 | |
| 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 | 375 | fun is_new_spass_version () = | 
| 
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 | 376 | string_ord (getenv "SPASS_VERSION", "3.7") = GREATER orelse | 
| 
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 | 377 | getenv "SPASS_NEW_HOME" <> "" | 
| 
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 | 378 | |
| 
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 | 379 | (* TODO: Legacy: Remove after planned Isabelle2012 release (and don't forget | 
| 
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 | 380 | "required_vars" and "script/spass"). *) | 
| 
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 | 381 | val spass_old_config : atp_config = | 
| 
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 | 382 |   {exec = (["ISABELLE_ATP"], "scripts/spass"),
 | 
| 
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 | 383 | required_vars = [["SPASS_OLD_HOME", "SPASS_HOME"]], | 
| 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 _ => | 
| 37962 
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
 blanchet parents: 
37926diff
changeset | 385 |      ("-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 | 386 | \-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 | 387 | |> sos = sosN ? prefix "-SOS=1 ", | 
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 388 |    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 | 389 | 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 | 390 | 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 | 391 | [(GaveUp, "SPASS beiseite: Completion found"), | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 392 | (TimedOut, "SPASS beiseite: Ran out of time"), | 
| 36965 | 393 | (OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"), | 
| 37413 | 394 | (MalformedInput, "Undefined symbol"), | 
| 37414 
d0cea0796295
expect SPASS 3.7, and give a friendly warning if an older version is used
 blanchet parents: 
37413diff
changeset | 395 | (MalformedInput, "Free Variable"), | 
| 44391 | 396 | (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 | 397 | (InternalError, "Please report this error")], | 
| 43466 
52f040bcfae7
tweaked TPTP formula kind for typing information used in the conjecture
 blanchet parents: 
43354diff
changeset | 398 | 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 | 399 | 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 | 400 | best_slices = fn ctxt => | 
| 42723 | 401 | (* FUDGE *) | 
| 46449 | 402 | [(0.333, (false, ((150, DFG DFG_Unsorted, "mono_tags??", liftingN, false), sosN))), | 
| 403 | (0.333, (false, ((300, DFG DFG_Unsorted, "poly_tags??", liftingN, false), sosN))), | |
| 404 | (0.334, (false, ((50, DFG DFG_Unsorted, "mono_tags??", liftingN, false), no_sosN)))] | |
| 405 | |> (if Config.get ctxt force_sos then hd #> apfst (K 1.0) #> single else I)} | |
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 406 | |
| 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 | 407 | val spass_old = (spass_oldN, spass_old_config) | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 408 | |
| 47149 | 409 | val spass_new_H1SOS = "-Heuristic=1 -SOS" | 
| 46449 | 410 | val spass_new_H2 = "-Heuristic=2" | 
| 46455 | 411 | val spass_new_H2SOS = "-Heuristic=2 -SOS" | 
| 412 | val spass_new_H2NuVS0 = "-Heuristic=2 -RNuV=1 -Sorts=0" | |
| 413 | val spass_new_H2NuVS0Red2 = | |
| 414 | "-Heuristic=2 -RNuV=1 -Sorts=0 -RFRew=2 -RBRew=2 -RTaut=2" | |
| 46449 | 415 | |
| 45301 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 416 | val spass_new_config : atp_config = | 
| 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 | 417 |   {exec = (["SPASS_NEW_HOME", "SPASS_HOME"], "SPASS"),
 | 
| 
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 | 418 | required_vars = [], | 
| 46444 | 419 | arguments = fn _ => fn _ => fn extra_options => fn timeout => fn _ => | 
| 46429 | 420 |      ("-Isabelle=1 -TimeLimit=" ^ string_of_int (to_secs 1 timeout))
 | 
| 46444 | 421 | |> extra_options <> "" ? prefix (extra_options ^ " "), | 
| 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 | 422 | proof_delims = #proof_delims spass_old_config, | 
| 
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 | 423 | known_failures = #known_failures spass_old_config, | 
| 
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 | 424 | conj_sym_kind = #conj_sym_kind spass_old_config, | 
| 
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 | 425 | prem_kind = #prem_kind spass_old_config, | 
| 46381 | 426 | best_slices = fn _ => | 
| 45301 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 427 | (* FUDGE *) | 
| 46455 | 428 | [(0.1667, (false, ((150, DFG DFG_Sorted, "mono_native", combsN, true), ""))), | 
| 429 | (0.1667, (false, ((500, DFG DFG_Sorted, "mono_native", liftingN, true), spass_new_H2SOS))), | |
| 430 | (0.1666, (false, ((50, DFG DFG_Sorted, "mono_native", liftingN, true), spass_new_H2))), | |
| 431 | (0.1000, (false, ((250, DFG DFG_Sorted, "mono_native", combsN, true), spass_new_H2NuVS0))), | |
| 432 | (0.1000, (false, ((100, DFG DFG_Sorted, "mono_native", combs_and_liftingN, true), spass_new_H2NuVS0))), | |
| 47149 | 433 | (0.1000, (false, ((1000, DFG DFG_Sorted, "mono_native", liftingN, true), spass_new_H1SOS))), | 
| 46455 | 434 | (0.1000, (false, ((150, DFG DFG_Sorted, "poly_guards??", liftingN, false), spass_new_H2NuVS0Red2))), | 
| 47149 | 435 | (0.1000, (false, ((300, DFG DFG_Sorted, "mono_native", combsN, true), spass_new_H2SOS)))]} | 
| 45301 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 436 | |
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 437 | val spass_new = (spass_newN, spass_new_config) | 
| 
866b075aa99b
added sorted DFG output for coming version of SPASS
 blanchet parents: 
45300diff
changeset | 438 | |
| 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 | 439 | fun spass () = | 
| 
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 | 440 | (spassN, if is_new_spass_version () then spass_new_config | 
| 
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 | 441 | else spass_old_config) | 
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 442 | |
| 37509 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 443 | (* Vampire *) | 
| 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 444 | |
| 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 | 445 | (* 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 | 446 | SystemOnTPTP. *) | 
| 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 | 447 | fun is_new_vampire_version () = | 
| 
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 | 448 | string_ord (getenv "VAMPIRE_VERSION", "1.8") = GREATER | 
| 44420 | 449 | |
| 44754 | 450 | val vampire_tff0 = TFF (TPTP_Monomorphic, TPTP_Implicit) | 
| 44589 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 blanchet parents: 
44586diff
changeset | 451 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 452 | val vampire_config : atp_config = | 
| 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 | 453 |   {exec = (["VAMPIRE_HOME"], "vampire"),
 | 
| 
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 | 454 | required_vars = [], | 
| 43569 
b342cd125533
removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
 blanchet parents: 
43567diff
changeset | 455 | arguments = fn _ => fn _ => fn sos => fn timeout => fn _ => | 
| 44417 | 456 | "--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 | 457 | " --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 | 458 | \ --forced_options propositional_to_bdd=off\ | 
| 44417 | 459 | \ --thanks \"Andrei and Krystof\" --input_file" | 
| 43473 
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
 blanchet parents: 
43467diff
changeset | 460 | |> 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 | 461 | proof_delims = | 
| 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 462 |      [("=========== Refutation ==========",
 | 
| 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 463 | "======= End of refutation ======="), | 
| 38033 | 464 |       ("% SZS output start Refutation", "% SZS output end Refutation"),
 | 
| 465 |       ("% 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 | 466 | known_failures = | 
| 45203 | 467 | 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 | 468 | [(GaveUp, "UNPROVABLE"), | 
| 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
42999diff
changeset | 469 | (GaveUp, "CANNOT PROVE"), | 
| 37509 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 blanchet parents: 
37506diff
changeset | 470 | (Unprovable, "Satisfiability detected"), | 
| 38647 
5500241da479
play with fudge factor + parse one more Vampire error
 blanchet parents: 
38646diff
changeset | 471 | (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 | 472 | (Interrupted, "Aborted by signal SIGINT")], | 
| 43466 
52f040bcfae7
tweaked TPTP formula kind for typing information used in the conjecture
 blanchet parents: 
43354diff
changeset | 473 | 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 | 474 | 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 | 475 | best_slices = fn ctxt => | 
| 42723 | 476 | (* FUDGE *) | 
| 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 | 477 | (if is_new_vampire_version () then | 
| 
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 | 478 | [(0.333, (false, ((150, vampire_tff0, "poly_guards??", combs_or_liftingN, false), sosN))), | 
| 
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 | 479 | (0.333, (false, ((500, vampire_tff0, "mono_native", combs_or_liftingN, false), sosN))), | 
| 
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 | 480 | (0.334, (true, ((50, vampire_tff0, "mono_native", combs_or_liftingN, false), no_sosN)))] | 
| 
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 | 481 | else | 
| 46449 | 482 | [(0.333, (false, ((150, FOF, "poly_guards??", combs_or_liftingN, false), sosN))), | 
| 483 | (0.333, (false, ((500, FOF, "mono_tags??", combs_or_liftingN, false), sosN))), | |
| 47055 
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
 blanchet parents: 
47053diff
changeset | 484 | (0.334, (true, ((50, FOF, "mono_guards??", combs_or_liftingN, false), no_sosN)))]) | 
| 44099 | 485 | |> (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 | 486 | else I)} | 
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 487 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 488 | 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 | 489 | |
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 490 | |
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 491 | (* Z3 with TPTP syntax *) | 
| 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 492 | |
| 44754 | 493 | val z3_tff0 = TFF (TPTP_Monomorphic, TPTP_Implicit) | 
| 44589 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 blanchet parents: 
44586diff
changeset | 494 | |
| 44423 
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
 blanchet parents: 
44422diff
changeset | 495 | val z3_tptp_config : atp_config = | 
| 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 | 496 |   {exec = (["Z3_HOME"], "z3"),
 | 
| 
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 | 497 | required_vars = [], | 
| 43354 
396aaa15dd8b
pass --trim option to "eproof" script to speed up proof reconstruction
 blanchet parents: 
43288diff
changeset | 498 | arguments = fn _ => fn _ => fn _ => fn timeout => fn _ => | 
| 44420 | 499 | "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 | 500 | proof_delims = [], | 
| 45203 | 501 | 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 | 502 | conj_sym_kind = Hypothesis, | 
| 
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
 blanchet parents: 
42708diff
changeset | 503 | prem_kind = Hypothesis, | 
| 42723 | 504 | best_slices = | 
| 44423 
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
 blanchet parents: 
44422diff
changeset | 505 | (* FUDGE *) | 
| 46435 | 506 | K [(0.5, (false, ((250, z3_tff0, "mono_native", combsN, false), ""))), | 
| 507 | (0.25, (false, ((125, z3_tff0, "mono_native", combsN, false), ""))), | |
| 508 | (0.125, (false, ((62, z3_tff0, "mono_native", combsN, false), ""))), | |
| 509 | (0.125, (false, ((31, z3_tff0, "mono_native", combsN, false), "")))]} | |
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 510 | |
| 44423 
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
 blanchet parents: 
44422diff
changeset | 511 | val z3_tptp = (z3_tptpN, z3_tptp_config) | 
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 512 | |
| 44590 | 513 | |
| 44754 | 514 | (* Not really a prover: Experimental Polymorphic TFF and THF output *) | 
| 44590 | 515 | |
| 44754 | 516 | fun dummy_config format type_enc : atp_config = | 
| 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 | 517 |   {exec = (["ISABELLE_ATP"], "scripts/dummy_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 | 518 | required_vars = [], | 
| 44590 | 519 | arguments = K (K (K (K (K "")))), | 
| 520 | proof_delims = [], | |
| 45203 | 521 | known_failures = known_szs_status_failures, | 
| 44590 | 522 | conj_sym_kind = Hypothesis, | 
| 523 | prem_kind = Hypothesis, | |
| 45521 | 524 | best_slices = | 
| 46407 
30e9720cc0b9
optimization: slice caching in case two consecutive slices are nearly identical
 blanchet parents: 
46402diff
changeset | 525 | K [(1.0, (false, ((200, format, type_enc, | 
| 
30e9720cc0b9
optimization: slice caching in case two consecutive slices are nearly identical
 blanchet parents: 
46402diff
changeset | 526 | if is_format_higher_order format then keep_lamsN | 
| 46409 
d4754183ccce
made option available to users (mostly for experiments)
 blanchet parents: 
46407diff
changeset | 527 | else combsN, false), "")))]} | 
| 44590 | 528 | |
| 45365 | 529 | val dummy_thf_format = THF (TPTP_Polymorphic, TPTP_Explicit, THF_With_Choice) | 
| 46435 | 530 | val dummy_thf_config = dummy_config dummy_thf_format "poly_native_higher" | 
| 45365 | 531 | val dummy_thf = (dummy_thfN, dummy_thf_config) | 
| 44754 | 532 | |
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 533 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 534 | (* Remote ATP invocation via SystemOnTPTP *) | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 535 | |
| 38061 
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
 blanchet parents: 
38049diff
changeset | 536 | val systems = Synchronized.var "atp_systems" ([] : string list) | 
| 31835 | 537 | |
| 538 | fun get_systems () = | |
| 44589 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 blanchet parents: 
44586diff
changeset | 539 | case Isabelle_System.bash_output | 
| 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 blanchet parents: 
44586diff
changeset | 540 | "\"$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 | 541 | (output, 0) => split_lines output | 
| 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39375diff
changeset | 542 | | (output, _) => | 
| 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39375diff
changeset | 543 | error (case extract_known_failure known_perl_failures output of | 
| 41744 | 544 | SOME failure => string_for_failure failure | 
| 39491 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39375diff
changeset | 545 | | NONE => perhaps (try (unsuffix "\n")) output ^ ".") | 
| 31835 | 546 | |
| 42537 
25ceb855a18b
improve version handling -- prefer versions of ToFoF, SInE, and SNARK that are known to work
 blanchet parents: 
42535diff
changeset | 547 | 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 | 548 | 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 | 549 | | 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 | 550 | 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 | 551 | 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 | 552 | | 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 | 553 | |
| 
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
 blanchet parents: 
38685diff
changeset | 554 | fun get_system name versions = | 
| 38589 
b03f8fe043ec
added "max_relevant_per_iter" option to Sledgehammer
 blanchet parents: 
38588diff
changeset | 555 | Synchronized.change_result systems | 
| 
b03f8fe043ec
added "max_relevant_per_iter" option to Sledgehammer
 blanchet parents: 
38588diff
changeset | 556 | (fn systems => (if null systems then get_systems () else systems) | 
| 42955 | 557 | |> `(`(find_system name versions))) | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 558 | |
| 38690 
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
 blanchet parents: 
38685diff
changeset | 559 | 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 | 560 | case get_system name versions of | 
| 42955 | 561 | (SOME sys, _) => sys | 
| 46480 | 562 |   | (NONE, []) => error ("SystemOnTPTP is not available.")
 | 
| 42955 | 563 | | (NONE, syss) => | 
| 46480 | 564 | case syss |> filter_out (String.isPrefix "%") | 
| 565 | |> filter_out (curry (op =) "") of | |
| 566 |       [] => error ("SystemOnTPTP is not available.")
 | |
| 567 |     | [msg] => error ("SystemOnTPTP is not available: " ^ msg ^ ".")
 | |
| 568 | | syss => | |
| 569 |       error ("System " ^ quote name ^ " is not available at SystemOnTPTP.\n" ^
 | |
| 570 | "(Available systems: " ^ commas_quote syss ^ ".)") | |
| 31835 | 571 | |
| 41148 | 572 | val max_remote_secs = 240 (* give Geoff Sutcliffe's servers a break *) | 
| 573 | ||
| 38690 
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
 blanchet parents: 
38685diff
changeset | 574 | 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 | 575 | conj_sym_kind prem_kind best_slice : atp_config = | 
| 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 | 576 |   {exec = (["ISABELLE_ATP"], "scripts/remote_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 | 577 | required_vars = [], | 
| 47074 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 578 | arguments = fn _ => fn _ => fn command => fn timeout => fn _ => | 
| 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 579 | (if command <> "" then "-c " ^ quote command ^ " " else "") ^ | 
| 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 580 | "-s " ^ the_system system_name system_versions ^ " " ^ | 
| 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 581 | "-t " ^ string_of_int (Int.min (max_remote_secs, to_secs 1 timeout)), | 
| 42962 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 blanchet parents: 
42955diff
changeset | 582 | proof_delims = union (op =) tstp_proof_delims proof_delims, | 
| 45203 | 583 | 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 | 584 | 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 | 585 | prem_kind = prem_kind, | 
| 47074 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 586 | best_slices = fn ctxt => [(1.0, (false, best_slice ctxt))]} | 
| 42443 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42332diff
changeset | 587 | |
| 43500 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 588 | 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 | 589 |         ({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 | 590 | : 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 | 591 | 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 | 592 | conj_sym_kind prem_kind best_slice | 
| 38023 | 593 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 594 | 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 | 595 | conj_sym_kind prem_kind best_slice = | 
| 40060 
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
 blanchet parents: 
40059diff
changeset | 596 | (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 | 597 | 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 | 598 | 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 | 599 | 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 | 600 | (remote_prefix ^ name, | 
| 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 601 | remotify_config system_name system_versions best_slice config) | 
| 28592 | 602 | |
| 44754 | 603 | val explicit_tff0 = TFF (TPTP_Monomorphic, TPTP_Explicit) | 
| 44589 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 blanchet parents: 
44586diff
changeset | 604 | |
| 43500 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 605 | val remote_e = | 
| 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 606 | remotify_atp e "EP" ["1.0", "1.1", "1.2"] | 
| 47074 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 607 | (K ((750, FOF, "mono_tags??", combsN, false), "") (* FUDGE *)) | 
| 44099 | 608 | val remote_leo2 = | 
| 609 | remotify_atp leo2 "LEO-II" ["1.2.8", "1.2.6"] | |
| 47074 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 610 | (K ((100, leo2_thf0, "mono_native_higher", liftingN, false), "") (* FUDGE *)) | 
| 44099 | 611 | val remote_satallax = | 
| 47074 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 612 | remotify_atp satallax "Satallax" ["2.3", "2.2", "2"] | 
| 47076 
f4838ce57772
removed Satallax option, now that this is the default
 blanchet parents: 
47074diff
changeset | 613 | (K ((100, satallax_thf0, "mono_native_higher", keep_lamsN, false), "") (* FUDGE *)) | 
| 43500 
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
 blanchet parents: 
43497diff
changeset | 614 | val remote_vampire = | 
| 44499 | 615 | remotify_atp vampire "Vampire" ["1.8"] | 
| 47074 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 616 | (K ((250, vampire_tff0, "mono_native", combs_or_liftingN, false), "") (* FUDGE *)) | 
| 44423 
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
 blanchet parents: 
44422diff
changeset | 617 | val remote_z3_tptp = | 
| 44754 | 618 | remotify_atp z3_tptp "Z3" ["3.0"] | 
| 47074 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 619 | (K ((250, z3_tff0, "mono_native", combsN, false), "") (* FUDGE *)) | 
| 44092 
bf489e54d7f8
renamed E wrappers for consistency with CASC conventions
 blanchet parents: 
43989diff
changeset | 620 | val remote_e_sine = | 
| 47074 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 621 | remote_atp e_sineN "SInE" ["0.4"] [] (#known_failures e_config) Axiom Conjecture | 
| 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 622 | (K ((500, FOF, "mono_guards??", combsN, false), "") (* FUDGE *)) | 
| 45338 | 623 | val remote_iprover = | 
| 624 | remote_atp iproverN "iProver" [] [] [] Axiom Conjecture | |
| 47074 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 625 | (K ((150, FOF, "mono_guards??", liftingN, false), "") (* FUDGE *)) | 
| 45338 | 626 | val remote_iprover_eq = | 
| 627 | remote_atp iprover_eqN "iProver-Eq" [] [] [] Axiom Conjecture | |
| 47074 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 628 | (K ((150, FOF, "mono_guards??", liftingN, false), "") (* FUDGE *)) | 
| 41740 
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
 blanchet parents: 
41738diff
changeset | 629 | val remote_snark = | 
| 42939 | 630 | remote_atp snarkN "SNARK" ["20080805r029", "20080805r024"] | 
| 45521 | 631 |       [("refutation.", "end_refutation.")] [] Hypothesis Hypothesis
 | 
| 47074 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 632 | (K ((100, explicit_tff0, "mono_native", liftingN, false), "") (* FUDGE *)) | 
| 44092 
bf489e54d7f8
renamed E wrappers for consistency with CASC conventions
 blanchet parents: 
43989diff
changeset | 633 | val remote_e_tofof = | 
| 44589 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 blanchet parents: 
44586diff
changeset | 634 | remote_atp e_tofofN "ToFoF" ["0.1"] [] (#known_failures e_config) Axiom | 
| 45521 | 635 | Hypothesis | 
| 47074 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 636 | (K ((150, explicit_tff0, "mono_native", liftingN, false), "") (* FUDGE *)) | 
| 42938 | 637 | val remote_waldmeister = | 
| 638 | remote_atp waldmeisterN "Waldmeister" ["710"] | |
| 45521 | 639 |       [("#START OF PROOF", "Proved Goals:")]
 | 
| 640 | [(OutOfResources, "Too many function symbols"), | |
| 641 | (Crashed, "Unrecoverable Segmentation Fault")] | |
| 642 | Hypothesis Hypothesis | |
| 47074 
101976132929
improve "remote_satallax" by exploiting unsat core
 blanchet parents: 
47073diff
changeset | 643 | (K ((50, CNF_UEQ, "mono_tags??", combsN, false), "") (* FUDGE *)) | 
| 38454 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 644 | |
| 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 645 | (* Setup *) | 
| 
9043eefe8d71
detect old Vampire and give a nicer error message
 blanchet parents: 
38433diff
changeset | 646 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 647 | fun add_atp (name, config) thy = | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 648 | 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 | 649 |   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 | 650 | |
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 651 | fun get_atp thy name = | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 652 | the (Symtab.lookup (Data.get thy) name) |> fst | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 653 |   handle Option.Option => error ("Unknown ATP: " ^ name ^ ".")
 | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 654 | |
| 41727 
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
 blanchet parents: 
41725diff
changeset | 655 | val supported_atps = Symtab.keys o Data.get | 
| 36371 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 656 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 657 | fun is_atp_installed thy name = | 
| 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 | 658 |   let val {exec, required_vars, ...} = get_atp thy name 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 | 659 | forall (exists (fn var => getenv var <> "")) (fst exec :: required_vars) | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 660 | end | 
| 36371 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 661 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 662 | fun refresh_systems_on_tptp () = | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 663 | Synchronized.change systems (fn _ => get_systems ()) | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39491diff
changeset | 664 | |
| 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 | 665 | 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 | 666 | 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 | 667 | if ord = smartN then | 
| 
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 | 668 | if atp = spass_newN orelse | 
| 
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 | 669 | (atp = spassN andalso is_new_spass_version ()) then | 
| 47073 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 blanchet parents: 
47055diff
changeset | 670 |         {is_lpo = false, gen_weights = true, gen_prec = true, gen_simp = false}
 | 
| 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 | 671 | 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 | 672 |         {is_lpo = false, gen_weights = false, gen_prec = false,
 | 
| 
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 | 673 | gen_simp = false} | 
| 
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 | 674 | 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 | 675 | let val is_lpo = String.isSubstring lpoN ord 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 | 676 |         {is_lpo = is_lpo,
 | 
| 
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 | 677 | gen_weights = not is_lpo andalso String.isSubstring xweightsN ord, | 
| 
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 | 678 | gen_prec = String.isSubstring xprecN ord, | 
| 
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 | 679 | gen_simp = String.isSubstring xsimpN ord} | 
| 
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 | 680 | 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 | 681 | 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 | 682 | |
| 
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 | 683 | fun atps () = | 
| 
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 | 684 | [alt_ergo, e, leo2, dummy_thf, satallax, spass_old, spass_new, spass (), | 
| 
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 | 685 | vampire, z3_tptp, remote_e, remote_e_sine, remote_e_tofof, remote_iprover, | 
| 
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 | 686 | remote_iprover_eq, remote_leo2, remote_satallax, remote_vampire, | 
| 
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 | 687 | remote_z3_tptp, remote_snark, remote_waldmeister] | 
| 
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 | 688 | |
| 
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 | 689 | fun setup thy = fold add_atp (atps ()) thy | 
| 35867 | 690 | |
| 28592 | 691 | end; |