# HG changeset patch # User blanchet # Date 1287991783 -7200 # Node ID e5ed638e49b00e7a03b4d4b8d44670ee0660c394 # Parent acb75271cdcefeccb9a08afab8f9ff625f83674c make "sledgehammer_params" work on single-threaded platforms diff -r acb75271cdce -r e5ed638e49b0 src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Fri Oct 22 18:31:45 2010 +0200 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Mon Oct 25 09:29:43 2010 +0200 @@ -157,13 +157,15 @@ prover :: avoid_too_many_local_threads thy n provers end +val num_processors = try Thread.numProcessors #> the_default 1 + (* 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 = let val thy = ProofContext.theory_of ctxt in [spassN, eN, vampireN, sine_eN (* FIXME: , smtN *)] |> map_filter (remotify_prover_if_not_installed ctxt) - |> avoid_too_many_local_threads thy (Thread.numProcessors ()) + |> avoid_too_many_local_threads thy (num_processors ()) |> space_implode " " end