use 30 s instead of 60 s as the default Sledgehammer timeout;
authorblanchet
Mon, 13 Sep 2010 13:12:33 +0200
changeset 39335 87a9ff4d5817
parent 39334 c0bb925ae912
child 39336 1899349a5026
use 30 s instead of 60 s as the default Sledgehammer timeout; very few proofs are lost this way (esp. thanks to the parallel use of provers, cf. Boehme & Nipkow 2010), and this is much more tolerable for users
doc-src/Sledgehammer/sledgehammer.tex
src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML
--- a/doc-src/Sledgehammer/sledgehammer.tex	Mon Sep 13 09:36:34 2010 +0200
+++ b/doc-src/Sledgehammer/sledgehammer.tex	Mon Sep 13 13:12:33 2010 +0200
@@ -434,7 +434,7 @@
 \opnodefault{atp}{string}
 Alias for \textit{atps}.
 
-\opdefault{timeout}{time}{$\mathbf{60}$ s}
+\opdefault{timeout}{time}{$\mathbf{30}$ s}
 Specifies the maximum amount of time that the ATPs should spend searching for a
 proof. For historical reasons, the default value of this option can be
 overridden using the option ``Sledgehammer: Time Limit'' from the ``Isabelle''
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Mon Sep 13 09:36:34 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Mon Sep 13 13:12:33 2010 +0200
@@ -50,7 +50,7 @@
 (*** parameters ***)
 
 val atps = Unsynchronized.ref ""
-val timeout = Unsynchronized.ref 60
+val timeout = Unsynchronized.ref 30
 val full_types = Unsynchronized.ref false
 
 val _ =