src/HOL/ex/Quickcheck_Narrowing_Examples.thy
changeset 42184 1d4fae76ba5e
parent 42024 51df23535105
child 43239 42f82fda796b
--- a/src/HOL/ex/Quickcheck_Narrowing_Examples.thy	Thu Mar 31 09:43:36 2011 +0200
+++ b/src/HOL/ex/Quickcheck_Narrowing_Examples.thy	Thu Mar 31 11:17:52 2011 +0200
@@ -32,15 +32,15 @@
 lemma "rev xs = xs"
   quickcheck[tester = narrowing, finite_types = false, default_type = int, expect = counterexample]
 oops
-
+(*
 lemma "rev xs = xs"
   quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
 oops
-
+*)
 subsection {* Simple examples with functions *}
 
 declare [[quickcheck_finite_functions]]
-
+(*
 lemma "map f xs = map g xs"
   quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
 oops
@@ -57,7 +57,7 @@
 lemma "map f xs = F f xs"
   quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
 oops
-
+*)
 lemma "map f xs = F f xs"
   quickcheck[tester = narrowing, finite_types = false, expect = counterexample]
 oops
@@ -163,7 +163,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 = 100, expect = counterexample]
+quickcheck[tester = narrowing, finite_types = false, default_type = nat, size = 6, timeout = 1000, expect = counterexample]
 oops