merged
authorbulwahn
Thu, 31 Mar 2011 17:15:13 +0200
changeset 42185 7101712baae8
parent 42184 1d4fae76ba5e (diff)
parent 42183 173b0f488428 (current diff)
child 42186 bb688200b949
child 42188 f6bc441fbf19
merged
--- a/src/HOL/Tools/Quickcheck/narrowing_generators.ML	Thu Mar 31 14:02:03 2011 +0200
+++ b/src/HOL/Tools/Quickcheck/narrowing_generators.ML	Thu Mar 31 17:15:13 2011 +0200
@@ -7,7 +7,7 @@
 signature NARROWING_GENERATORS =
 sig
   val compile_generator_expr:
-    Proof.context -> term * term list -> int list -> term list option * Quickcheck.report option
+    Proof.context -> (term * term list) list -> int list -> term list option * Quickcheck.report option
   val put_counterexample: (unit -> term list option) -> Proof.context -> Proof.context
   val finite_functions : bool Config.T
   val setup: theory -> theory
@@ -212,7 +212,7 @@
     list_abs (names ~~ Ts', t'')
   end
 
-fun compile_generator_expr ctxt (t, eval_terms) [size] =
+fun compile_generator_expr ctxt [(t, eval_terms)] [_, size] =
   let
     val thy = ProofContext.theory_of ctxt
     val t' = list_abs_free (Term.add_frees t [], t)
@@ -221,7 +221,7 @@
       Const (@{const_name Quickcheck_Narrowing.ensure_testable}, fastype_of t --> fastype_of t) $ t
     val result = dynamic_value_strict
       (Counterexample.get, Counterexample.put, "Narrowing_Generators.put_counterexample")
-      thy (Option.map o map) (ensure_testable t'') []
+      thy (Option.map o map) (ensure_testable t'') [] size
   in
     (result, NONE)
   end;
--- a/src/HOL/ex/Quickcheck_Narrowing_Examples.thy	Thu Mar 31 14:02:03 2011 +0200
+++ b/src/HOL/ex/Quickcheck_Narrowing_Examples.thy	Thu Mar 31 17:15:13 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