--- a/src/HOL/Tools/ATP/atp_systems.ML Thu Feb 02 10:16:10 2012 +0100
+++ b/src/HOL/Tools/ATP/atp_systems.ML Thu Feb 02 12:42:05 2012 +0100
@@ -352,7 +352,8 @@
(* Experimental *)
val spass_new_config : atp_config =
{exec = ("ISABELLE_ATP", "scripts/spass_new"),
- required_execs = [],
+ required_execs =
+ [("SPASS_NEW_HOME", "SPASS"), ("SPASS_NEW_HOME", "tptp2dfg")],
arguments = fn _ => fn _ => fn incomplete => fn timeout => fn _ =>
("-Auto -LR=1 -Isabelle=1 -TimeLimit=" ^ string_of_int (to_secs 1 timeout))
|> incomplete = spass_incompleteN ? prefix "-Splits=0 -FullRed=0 ",
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Thu Feb 02 10:16:10 2012 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Thu Feb 02 12:42:05 2012 +0100
@@ -210,7 +210,8 @@
(* The first ATP of the list is used by Auto Sledgehammer. Because of the low
timeout, it makes sense to put SPASS first. *)
fun default_provers_param_value ctxt =
- [spassN, eN, vampireN, SMT_Solver.solver_name_of ctxt, e_sineN, waldmeisterN]
+ [spassN, spass_newN, eN, vampireN, SMT_Solver.solver_name_of ctxt, e_sineN,
+ waldmeisterN]
|> map_filter (remotify_prover_if_not_installed ctxt)
|> avoid_too_many_threads ctxt (Multithreading.max_threads_value (),
max_default_remote_threads)