--- a/src/HOL/Tools/Quickcheck/exhaustive_generators.ML Wed Nov 09 11:34:59 2011 +0100
+++ b/src/HOL/Tools/Quickcheck/exhaustive_generators.ML Wed Nov 09 11:35:09 2011 +0100
@@ -488,7 +488,7 @@
thy (SOME target) (K I) (HOLogic.mk_list @{typ "code_numeral => bool"} ts') []
end;
-val test_goals = Quickcheck_Common.generator_test_goal_terms compile_generator_expr;
+val test_goals = Quickcheck_Common.generator_test_goal_terms ("exhaustive", compile_generator_expr);
(* setup *)
--- a/src/HOL/Tools/Quickcheck/narrowing_generators.ML Wed Nov 09 11:34:59 2011 +0100
+++ b/src/HOL/Tools/Quickcheck/narrowing_generators.ML Wed Nov 09 11:35:09 2011 +0100
@@ -263,7 +263,7 @@
(NONE, !current_result)
else
let
- val _ = message ("[Quickcheck-Narrowing] Test data size: " ^ string_of_int k)
+ val _ = message ("[Quickcheck-narrowing] Test data size: " ^ string_of_int k)
val _ = current_size := k
val ((response, _), timing) = elapsed_time ("execution of size " ^ string_of_int k)
(fn () => Isabelle_System.bash_output (executable ^ " " ^ string_of_int k))
--- a/src/HOL/Tools/Quickcheck/quickcheck_common.ML Wed Nov 09 11:34:59 2011 +0100
+++ b/src/HOL/Tools/Quickcheck/quickcheck_common.ML Wed Nov 09 11:35:09 2011 +0100
@@ -16,7 +16,8 @@
val collect_results : ('a -> Quickcheck.result) -> 'a list -> Quickcheck.result list -> Quickcheck.result list
type compile_generator =
Proof.context -> (term * term list) list -> int list -> term list option * Quickcheck.report option
- val generator_test_goal_terms : compile_generator -> Proof.context -> bool -> (string * typ) list
+ val generator_test_goal_terms :
+ string * compile_generator -> Proof.context -> bool -> (string * typ) list
-> (term * term list) list -> Quickcheck.result list
val ensure_sort_datatype:
((sort * sort) * sort) * (Datatype.config -> Datatype.descr -> (string * sort) list
@@ -27,7 +28,8 @@
-> Proof.context -> (term * term list) list -> term
val mk_fun_upd : typ -> typ -> term * term -> term -> term
val post_process_term : term -> term
- val test_term : compile_generator -> Proof.context -> bool -> term * term list -> Quickcheck.result
+ val test_term : string * compile_generator
+ -> Proof.context -> bool -> term * term list -> Quickcheck.result
end;
structure Quickcheck_Common : QUICKCHECK_COMMON =
@@ -61,7 +63,7 @@
let val ({cpu, ...}, result) = Timing.timing e ()
in (result, (description, Time.toMilliseconds cpu)) end
-fun test_term compile ctxt catch_code_errors (t, eval_terms) =
+fun test_term (name, compile) ctxt catch_code_errors (t, eval_terms) =
let
val _ = check_test_term t
val names = Term.add_free_names t []
@@ -74,7 +76,8 @@
NONE
else
let
- val _ = Quickcheck.message ctxt ("Test data size: " ^ string_of_int k)
+ val _ = Quickcheck.message ctxt
+ ("[Quickcheck-" ^ name ^ "]Test data size: " ^ string_of_int k)
val _ = current_size := k
val ((result, report), timing) =
cpu_time ("size " ^ string_of_int k) (fn () => test_fun [1, k - 1])
@@ -89,7 +92,8 @@
val _ = Quickcheck.add_timing comp_time current_result
in
case test_fun of
- NONE => (Quickcheck.message ctxt "Conjecture is not executable with quickcheck"; !current_result)
+ NONE => (Quickcheck.message ctxt ("Conjecture is not executable with Quickcheck-" ^ name);
+ !current_result)
| SOME test_fun =>
let
val (response, exec_time) =
@@ -109,7 +113,8 @@
if k > size then true
else if tester k then with_size tester (k + 1) else false
val (results, exec_time) = cpu_time "quickcheck batch execution" (fn () =>
- Option.map (map (fn test_fun => TimeLimit.timeLimit (seconds (Config.get ctxt Quickcheck.timeout))
+ Option.map (map (fn test_fun =>
+ TimeLimit.timeLimit (seconds (Config.get ctxt Quickcheck.timeout))
(fn () => with_size test_fun 1) ()
handle TimeLimit.TimeOut => true)) test_funs)
in
@@ -127,7 +132,8 @@
if k > size then NONE
else case tester k of SOME ts => SOME ts | NONE => with_size tester (k + 1)
val (results, exec_time) = cpu_time "quickcheck batch execution" (fn () =>
- Option.map (map (fn test_fun => TimeLimit.timeLimit (seconds (Config.get ctxt Quickcheck.timeout))
+ Option.map (map (fn test_fun =>
+ TimeLimit.timeLimit (seconds (Config.get ctxt Quickcheck.timeout))
(fn () => with_size test_fun 1) ()
handle TimeLimit.TimeOut => NONE)) test_funs)
in
@@ -136,7 +142,7 @@
end
-fun test_term_with_increasing_cardinality compile ctxt catch_code_errors ts =
+fun test_term_with_cardinality (name, compile) ctxt catch_code_errors ts =
let
val thy = Proof_Context.theory_of ctxt
val (ts', eval_terms) = split_list ts
@@ -167,7 +173,8 @@
val _ = Quickcheck.add_timing comp_time current_result
in
case test_fun of
- NONE => (Quickcheck.message ctxt "Conjecture is not executable with quickcheck"; !current_result)
+ NONE => (Quickcheck.message ctxt ("Conjecture is not executable with Quickcheck-" ^ name);
+ !current_result)
| SOME test_fun =>
let
val _ = case get_first (test_card_size test_fun) enumeration_card_size of
@@ -207,14 +214,14 @@
fun map_goal_and_eval_terms f (check_goal, eval_terms) = (f check_goal, map f eval_terms)
val thy = Proof_Context.theory_of lthy
val default_insts =
- if Config.get lthy Quickcheck.finite_types then (get_finite_types lthy) else (Quickcheck.default_type lthy)
+ if Config.get lthy Quickcheck.finite_types then get_finite_types else Quickcheck.default_type
val inst_goals =
map (fn (check_goal, eval_terms) =>
if not (null (Term.add_tfree_names check_goal [])) then
map (fn T =>
(pair (SOME T) o Term o apfst (Object_Logic.atomize_term thy))
(map_goal_and_eval_terms (monomorphic_term thy insts T) (check_goal, eval_terms))
- handle WELLSORTED s => (SOME T, Wellsorted_Error s)) default_insts
+ handle WELLSORTED s => (SOME T, Wellsorted_Error s)) (default_insts lthy)
else
[(NONE, Term (Object_Logic.atomize_term thy check_goal, eval_terms))]) goals
val error_msg =
@@ -247,19 +254,19 @@
SOME (lhs, rhs) => (t, eval_terms @ [rhs, lhs])
| NONE => (t, eval_terms)
-fun generator_test_goal_terms compile ctxt catch_code_errors insts goals =
+fun generator_test_goal_terms (name, compile) ctxt catch_code_errors insts goals =
let
fun test_term' goal =
case goal of
- [(NONE, t)] => test_term compile ctxt catch_code_errors t
- | ts => test_term_with_increasing_cardinality compile ctxt catch_code_errors (map snd ts)
+ [(NONE, t)] => test_term (name, compile) ctxt catch_code_errors t
+ | ts => test_term_with_cardinality (name, compile) ctxt catch_code_errors (map snd ts)
val goals' = instantiate_goals ctxt insts goals
|> map (map (apsnd add_equation_eval_terms))
in
if Config.get ctxt Quickcheck.finite_types then
collect_results test_term' goals' []
else
- collect_results (test_term compile ctxt catch_code_errors)
+ collect_results (test_term (name, compile) ctxt catch_code_errors)
(maps (map snd) goals') []
end;
@@ -411,6 +418,5 @@
| NONE => process_args t)
| _ => process_args t
end
-
end;
--- a/src/HOL/Tools/Quickcheck/random_generators.ML Wed Nov 09 11:34:59 2011 +0100
+++ b/src/HOL/Tools/Quickcheck/random_generators.ML Wed Nov 09 11:35:09 2011 +0100
@@ -438,7 +438,7 @@
end
end;
-val test_goals = Quickcheck_Common.generator_test_goal_terms compile_generator_expr;
+val test_goals = Quickcheck_Common.generator_test_goal_terms ("random", compile_generator_expr);
(** setup **)