increasing quickcheck's timeout in the example theory to avoid failures on the testing infrastructure
authorbulwahn
Tue, 06 Dec 2011 15:23:16 +0100
changeset 45773 7f2366dc8c0f
parent 45772 8a8f78ce0dcf
child 45774 9b11989f38b0
increasing quickcheck's timeout in the example theory to avoid failures on the testing infrastructure
src/HOL/ex/Quickcheck_Narrowing_Examples.thy
--- a/src/HOL/ex/Quickcheck_Narrowing_Examples.thy	Tue Dec 06 14:29:37 2011 +0100
+++ b/src/HOL/ex/Quickcheck_Narrowing_Examples.thy	Tue Dec 06 15:23:16 2011 +0100
@@ -9,6 +9,8 @@
 imports Main
 begin
 
+declare [[quickcheck_timeout = 3600]]
+
 subsection {* Minimalistic examples *}
 
 lemma
@@ -98,7 +100,7 @@
 
 lemma
   "list_all2 P (rev xs) (rev ys) = list_all2 P xs (rev ys)"
-  quickcheck[tester = narrowing, finite_types = false, expect = counterexample, timeout = 60]
+  quickcheck[tester = narrowing, finite_types = false, expect = counterexample]
 oops
 
 lemma "map f xs = F f xs"
@@ -224,7 +226,7 @@
 
 lemma is_ord_l_bal:
  "\<lbrakk> is_ord(MKT (x :: nat) l r h); height l = height r + 2 \<rbrakk> \<Longrightarrow> is_ord(l_bal(x,l,r))"
-quickcheck[tester = narrowing, finite_types = false, default_type = nat, size = 6, timeout = 1000, expect = counterexample]
+quickcheck[tester = narrowing, finite_types = false, default_type = nat, size = 6, expect = counterexample]
 oops