author | bulwahn |
Wed, 21 Jul 2010 19:21:07 +0200 | |
changeset 37929 | 22e0797857e6 |
parent 37921 | 1e846be00ddf |
child 37930 | 38e71ffc8fe8 |
--- a/src/HOL/ex/Quickcheck_Examples.thy Wed Jul 21 18:13:15 2010 +0200 +++ b/src/HOL/ex/Quickcheck_Examples.thy Wed Jul 21 19:21:07 2010 +0200 @@ -19,27 +19,27 @@ subsection {* Lists *} theorem "map g (map f xs) = map (g o f) xs" - quickcheck + quickcheck[expect = no_counterexample] oops theorem "map g (map f xs) = map (f o g) xs" - quickcheck + quickcheck[expect = counterexample] oops theorem "rev (xs @ ys) = rev ys @ rev xs" - quickcheck + quickcheck[expect = no_counterexample] oops theorem "rev (xs @ ys) = rev xs @ rev ys" - quickcheck + quickcheck[expect = counterexample] oops theorem "rev (rev xs) = xs" - quickcheck + quickcheck[expect = no_counterexample] oops theorem "rev xs = xs" - quickcheck + quickcheck[expect = counterexample] oops text {* An example involving functions inside other data structures *} @@ -49,11 +49,11 @@ | "app (f # fs) x = app fs (f x)" lemma "app (fs @ gs) x = app gs (app fs x)" - quickcheck + quickcheck[expect = no_counterexample] by (induct fs arbitrary: x) simp_all lemma "app (fs @ gs) x = app fs (app gs x)" - quickcheck + quickcheck[expect = counterexample] oops primrec occurs :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where @@ -67,16 +67,16 @@ text {* A lemma, you'd think to be true from our experience with delAll *} lemma "Suc (occurs a (del1 a xs)) = occurs a xs" -- {* Wrong. Precondition needed.*} - quickcheck + quickcheck[expect = counterexample] oops lemma "xs ~= [] \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs" - quickcheck + quickcheck[expect = counterexample] -- {* Also wrong.*} oops lemma "0 < occurs a xs \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs" - quickcheck + quickcheck[expect = no_counterexample] by (induct xs) auto primrec replace :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where @@ -85,12 +85,12 @@ else (x#(replace a b xs)))" lemma "occurs a xs = occurs b (replace a b xs)" - quickcheck + quickcheck[expect = counterexample] -- {* Wrong. Precondition needed.*} oops lemma "occurs b xs = 0 \<or> a=b \<longrightarrow> occurs a xs = occurs b (replace a b xs)" - quickcheck + quickcheck[expect = no_counterexample] by (induct xs) simp_all @@ -113,12 +113,12 @@ | "mirror (Branch l r) = Branch (mirror r) (mirror l)" theorem "plant (rev (leaves xt)) = mirror xt" - quickcheck + quickcheck[expect = counterexample] --{* Wrong! *} oops theorem "plant((leaves xt) @ (leaves yt)) = Branch xt yt" - quickcheck + quickcheck[expect = counterexample] --{* Wrong! *} oops @@ -133,7 +133,7 @@ | "root (Node f x y) = f" theorem "hd (inOrder xt) = root xt" - quickcheck + quickcheck[expect = counterexample] --{* Wrong! *} oops
--- a/src/HOL/ex/Quickcheck_Lattice_Examples.thy Wed Jul 21 18:13:15 2010 +0200 +++ b/src/HOL/ex/Quickcheck_Lattice_Examples.thy Wed Jul 21 19:21:07 2010 +0200 @@ -22,109 +22,109 @@ lemma sup_inf_distrib2: "((y :: 'a :: distrib_lattice) \<sqinter> z) \<squnion> x = (y \<squnion> x) \<sqinter> (z \<squnion> x)" - quickcheck + quickcheck[expect = no_counterexample] by(simp add: inf_sup_aci sup_inf_distrib1) lemma sup_inf_distrib2_1: "((y :: 'a :: lattice) \<sqinter> z) \<squnion> x = (y \<squnion> x) \<sqinter> (z \<squnion> x)" - quickcheck + quickcheck[expect = counterexample] oops lemma sup_inf_distrib2_2: "((y :: 'a :: distrib_lattice) \<sqinter> z') \<squnion> x = (y \<squnion> x) \<sqinter> (z \<squnion> x)" - quickcheck + quickcheck[expect = counterexample] oops lemma inf_sup_distrib1_1: "(x :: 'a :: distrib_lattice) \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x' \<sqinter> z)" - quickcheck + quickcheck[expect = counterexample] oops lemma inf_sup_distrib2_1: "((y :: 'a :: distrib_lattice) \<squnion> z) \<sqinter> x = (y \<sqinter> x) \<squnion> (y \<sqinter> x)" - quickcheck + quickcheck[expect = counterexample] oops subsection {* Bounded lattices *} lemma inf_bot_left [simp]: "\<bottom> \<sqinter> (x :: 'a :: bounded_lattice_bot) = \<bottom>" - quickcheck + quickcheck[expect = no_counterexample] by (rule inf_absorb1) simp lemma inf_bot_left_1: "\<bottom> \<sqinter> (x :: 'a :: bounded_lattice_bot) = x" - quickcheck + quickcheck[expect = counterexample] oops lemma inf_bot_left_2: "y \<sqinter> (x :: 'a :: bounded_lattice_bot) = \<bottom>" - quickcheck + quickcheck[expect = counterexample] oops lemma inf_bot_left_3: "x \<noteq> \<bottom> ==> y \<sqinter> (x :: 'a :: bounded_lattice_bot) \<noteq> \<bottom>" - quickcheck + quickcheck[expect = counterexample] oops lemma inf_bot_right [simp]: "(x :: 'a :: bounded_lattice_bot) \<sqinter> \<bottom> = \<bottom>" - quickcheck + quickcheck[expect = no_counterexample] by (rule inf_absorb2) simp lemma inf_bot_right_1: "x \<noteq> \<bottom> ==> (x :: 'a :: bounded_lattice_bot) \<sqinter> \<bottom> = y" - quickcheck + quickcheck[expect = counterexample] oops lemma inf_bot_right_2: "(x :: 'a :: bounded_lattice_bot) \<sqinter> \<bottom> ~= \<bottom>" - quickcheck + quickcheck[expect = counterexample] oops lemma inf_bot_right [simp]: "(x :: 'a :: bounded_lattice_bot) \<squnion> \<bottom> = \<bottom>" - quickcheck + quickcheck[expect = counterexample] oops lemma sup_bot_left [simp]: "\<bottom> \<squnion> (x :: 'a :: bounded_lattice_bot) = x" - quickcheck + quickcheck[expect = no_counterexample] by (rule sup_absorb2) simp lemma sup_bot_right [simp]: "(x :: 'a :: bounded_lattice_bot) \<squnion> \<bottom> = x" - quickcheck + quickcheck[expect = no_counterexample] by (rule sup_absorb1) simp lemma sup_eq_bot_iff [simp]: "(x :: 'a :: bounded_lattice_bot) \<squnion> y = \<bottom> \<longleftrightarrow> x = \<bottom> \<and> y = \<bottom>" - quickcheck + quickcheck[expect = no_counterexample] by (simp add: eq_iff) lemma sup_top_left [simp]: "\<top> \<squnion> (x :: 'a :: bounded_lattice_top) = \<top>" - quickcheck + quickcheck[expect = no_counterexample] by (rule sup_absorb1) simp lemma sup_top_right [simp]: "(x :: 'a :: bounded_lattice_top) \<squnion> \<top> = \<top>" - quickcheck + quickcheck[expect = no_counterexample] by (rule sup_absorb2) simp lemma inf_top_left [simp]: "\<top> \<sqinter> x = (x :: 'a :: bounded_lattice_top)" - quickcheck + quickcheck[expect = no_counterexample] by (rule inf_absorb2) simp lemma inf_top_right [simp]: "x \<sqinter> \<top> = (x :: 'a :: bounded_lattice_top)" - quickcheck + quickcheck[expect = no_counterexample] by (rule inf_absorb1) simp lemma inf_eq_top_iff [simp]: "x \<sqinter> y = \<top> \<longleftrightarrow> x = \<top> \<and> y = \<top>" - quickcheck + quickcheck[expect = no_counterexample] by (simp add: eq_iff)
--- a/src/Tools/quickcheck.ML Wed Jul 21 18:13:15 2010 +0200 +++ b/src/Tools/quickcheck.ML Wed Jul 21 19:21:07 2010 +0200 @@ -13,8 +13,10 @@ datatype report = Report of { iterations : int, raised_match_errors : int, satisfied_assms : int list, positive_concl_tests : int } + datatype expectation = No_Expectation | No_Counterexample | Counterexample; datatype test_params = Test_Params of - { size: int, iterations: int, default_type: typ list, no_assms: bool, report: bool, quiet : bool}; + { size: int, iterations: int, default_type: typ list, no_assms: bool, + expect : expectation, report: bool, quiet : bool}; val test_params_of: theory -> test_params val add_generator: string * (Proof.context -> bool -> term -> int -> term list option * (bool list * bool)) @@ -65,32 +67,41 @@ (if null satisfied_assms then replicate (length assms) 0 else satisfied_assms), positive_concl_tests = if concl then positive_concl_tests + 1 else positive_concl_tests} +(* expectation *) + +datatype expectation = No_Expectation | No_Counterexample | Counterexample; + +fun merge_expectation (expect1, expect2) = + if expect1 = expect2 then expect1 else No_Expectation + (* quickcheck configuration -- default parameters, test generators *) datatype test_params = Test_Params of - { size: int, iterations: int, default_type: typ list, no_assms: bool, report: bool, quiet : bool}; + { size: int, iterations: int, default_type: typ list, no_assms: bool, + expect : expectation, report: bool, quiet : bool}; -fun dest_test_params (Test_Params { size, iterations, default_type, no_assms, report, quiet }) = - ((size, iterations), ((default_type, no_assms), (report, quiet))); -fun make_test_params ((size, iterations), ((default_type, no_assms), (report, quiet))) = +fun dest_test_params (Test_Params { size, iterations, default_type, no_assms, expect, report, quiet }) = + ((size, iterations), ((default_type, no_assms), ((expect, report), quiet))); +fun make_test_params ((size, iterations), ((default_type, no_assms), ((expect, report), quiet))) = Test_Params { size = size, iterations = iterations, default_type = default_type, - no_assms = no_assms, report = report, quiet = quiet }; -fun map_test_params f (Test_Params { size, iterations, default_type, no_assms, report, quiet }) = - make_test_params (f ((size, iterations), ((default_type, no_assms), (report, quiet)))); + no_assms = no_assms, expect = expect, report = report, quiet = quiet }; +fun map_test_params f (Test_Params { size, iterations, default_type, no_assms, expect, report, quiet }) = + make_test_params (f ((size, iterations), ((default_type, no_assms), ((expect, report), quiet)))); fun merge_test_params (Test_Params { size = size1, iterations = iterations1, default_type = default_type1, - no_assms = no_assms1, report = report1, quiet = quiet1 }, + no_assms = no_assms1, expect = expect1, report = report1, quiet = quiet1 }, Test_Params { size = size2, iterations = iterations2, default_type = default_type2, - no_assms = no_assms2, report = report2, quiet = quiet2 }) = + no_assms = no_assms2, expect = expect2, report = report2, quiet = quiet2 }) = make_test_params ((Int.max (size1, size2), Int.max (iterations1, iterations2)), ((merge (op =) (default_type1, default_type2), no_assms1 orelse no_assms2), - (report1 orelse report2, quiet1 orelse quiet2))); + ((merge_expectation (expect1, expect2), report1 orelse report2), quiet1 orelse quiet2))); structure Data = Theory_Data ( type T = (string * (Proof.context -> bool -> term -> int -> term list option * (bool list * bool))) list * test_params; val empty = ([], Test_Params - { size = 10, iterations = 100, default_type = [], no_assms = false, report = false, quiet = false}); + { size = 10, iterations = 100, default_type = [], no_assms = false, + expect = No_Expectation, report = false, quiet = false}); val extend = I; fun merge ((generators1, params1), (generators2, params2)) : T = (AList.merge (op =) (K true) (generators1, generators2), @@ -271,7 +282,7 @@ let val ctxt = Proof.context_of state; val assms = map term_of (Assumption.all_assms_of ctxt); - val Test_Params { size, iterations, default_type, no_assms, report, quiet } = + val Test_Params {size, iterations, default_type, no_assms, ...} = (snd o Data.get o Proof.theory_of) state; val res = try (test_goal true false NONE size iterations default_type no_assms [] 1 assms) state; @@ -297,6 +308,11 @@ | read_bool "true" = true | read_bool s = error ("Not a Boolean value: " ^ s) +fun read_expectation "no_expectation" = No_Expectation + | read_expectation "no_counterexample" = No_Counterexample + | read_expectation "counterexample" = Counterexample + | read_expectation s = error ("Not an expectation value: " ^ s) + fun parse_test_param ctxt ("size", [arg]) = (apfst o apfst o K) (read_nat arg) | parse_test_param ctxt ("iterations", [arg]) = @@ -305,10 +321,12 @@ (apsnd o apfst o apfst o K) (map (ProofContext.read_typ ctxt) arg) | parse_test_param ctxt ("no_assms", [arg]) = (apsnd o apfst o apsnd o K) (read_bool arg) + | parse_test_param ctxt ("expect", [arg]) = + (apsnd o apsnd o apfst o apfst o K) (read_expectation arg) | parse_test_param ctxt ("report", [arg]) = - (apsnd o apsnd o apfst o K) (read_bool arg) + (apsnd o apsnd o apfst o apsnd o K) (read_bool arg) | parse_test_param ctxt ("quiet", [arg]) = - (apsnd o apsnd o apsnd o K) (read_bool arg) + (apsnd o apsnd o apsnd o K) (read_bool arg) | parse_test_param ctxt (name, _) = error ("Unknown test parameter: " ^ name); @@ -336,10 +354,14 @@ val assms = map term_of (Assumption.all_assms_of ctxt); val default_params = (dest_test_params o snd o Data.get) thy; val f = fold (parse_test_param_inst ctxt) args; - val (((size, iterations), ((default_type, no_assms), (report, quiet))), (generator_name, insts)) = + val (((size, iterations), ((default_type, no_assms), ((expect, report), quiet))), (generator_name, insts)) = f (default_params, (NONE, [])); in test_goal quiet report generator_name size iterations default_type no_assms insts i assms state + |> tap (fn (SOME x, _) => if expect = No_Counterexample then + error ("quickcheck expect to find no counterexample but found one") else () + | (NONE, _) => if expect = Counterexample then + error ("quickcheck expected to find a counterexample but did not find one") else ()) end; fun quickcheck args i state = fst (gen_quickcheck args i state);