| author | wenzelm | 
| Mon, 24 May 2010 23:01:51 +0200 | |
| changeset 37072 | 9105c8237c7a | 
| parent 36960 | 01594f816e3a | 
| child 37909 | 583543ad6ad1 | 
| permissions | -rw-r--r-- | 
| 30824 | 1 | (* Title: Tools/quickcheck.ML | 
| 28256 | 2 | Author: Stefan Berghofer, Florian Haftmann, TU Muenchen | 
| 3 | ||
| 4 | Generic counterexample search engine. | |
| 5 | *) | |
| 6 | ||
| 7 | signature QUICKCHECK = | |
| 8 | sig | |
| 32740 | 9 | val auto: bool Unsynchronized.ref | 
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34128diff
changeset | 10 | val timing : bool Unsynchronized.ref | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 11 | datatype report = Report of | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 12 |     { iterations : int, raised_match_errors : int,
 | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 13 | satisfied_assms : int list, positive_concl_tests : int } | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 14 | val gen_test_term: Proof.context -> bool -> bool -> string option -> int -> int -> term -> | 
| 35380 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35379diff
changeset | 15 | (string * term) list option * ((string * int) list * ((int * report list) list) option) | 
| 30980 | 16 | val test_term: Proof.context -> bool -> string option -> int -> int -> term -> | 
| 17 | (string * term) list option | |
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 18 | val add_generator: | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 19 | string * (Proof.context -> bool -> term -> int -> term list option * (bool list * bool)) | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 20 | -> theory -> theory | 
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 21 | val setup: theory -> theory | 
| 32297 | 22 | val quickcheck: (string * string) list -> int -> Proof.state -> (string * term) list option | 
| 28256 | 23 | end; | 
| 24 | ||
| 25 | structure Quickcheck : QUICKCHECK = | |
| 26 | struct | |
| 27 | ||
| 30980 | 28 | (* preferences *) | 
| 29 | ||
| 32740 | 30 | val auto = Unsynchronized.ref false; | 
| 30980 | 31 | |
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34128diff
changeset | 32 | val timing = Unsynchronized.ref false; | 
| 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34128diff
changeset | 33 | |
| 30980 | 34 | val _ = | 
| 35 | ProofGeneralPgip.add_preference Preferences.category_tracing | |
| 32966 
5b21661fe618
indicate CRITICAL nature of various setmp combinators;
 wenzelm parents: 
32859diff
changeset | 36 | (setmp_CRITICAL auto true (fn () => | 
| 30980 | 37 | Preferences.bool_pref auto | 
| 38 | "auto-quickcheck" | |
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 39 | "Whether to run Quickcheck automatically.") ()); | 
| 30980 | 40 | |
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 41 | (* quickcheck report *) | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 42 | |
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 43 | datatype single_report = Run of bool list * bool | MatchExc | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 44 | |
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 45 | datatype report = Report of | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 46 |   { iterations : int, raised_match_errors : int,
 | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 47 | satisfied_assms : int list, positive_concl_tests : int } | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 48 | |
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 49 | fun collect_single_report single_report | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 50 |     (Report {iterations = iterations, raised_match_errors = raised_match_errors,
 | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 51 | satisfied_assms = satisfied_assms, positive_concl_tests = positive_concl_tests}) = | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 52 | case single_report | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 53 | of MatchExc => | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 54 |     Report {iterations = iterations + 1, raised_match_errors = raised_match_errors + 1,
 | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 55 | satisfied_assms = satisfied_assms, positive_concl_tests = positive_concl_tests} | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 56 | | Run (assms, concl) => | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 57 |     Report {iterations = iterations + 1, raised_match_errors = raised_match_errors,
 | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 58 | satisfied_assms = | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 59 | map2 (fn b => fn s => if b then s + 1 else s) assms | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 60 | (if null satisfied_assms then replicate (length assms) 0 else satisfied_assms), | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 61 | positive_concl_tests = if concl then positive_concl_tests + 1 else positive_concl_tests} | 
| 30973 
304ab57afa6e
observe distinction between Pure/Tools and Tools more closely
 haftmann parents: 
30824diff
changeset | 62 | |
| 28315 | 63 | (* quickcheck configuration -- default parameters, test generators *) | 
| 64 | ||
| 28309 | 65 | datatype test_params = Test_Params of | 
| 35379 | 66 |   { size: int, iterations: int, default_type: typ option, no_assms: bool, report: bool, quiet : bool};
 | 
| 28309 | 67 | |
| 35379 | 68 | fun dest_test_params (Test_Params { size, iterations, default_type, no_assms, report, quiet }) =
 | 
| 69 | ((size, iterations), ((default_type, no_assms), (report, quiet))); | |
| 70 | fun make_test_params ((size, iterations), ((default_type, no_assms), (report, quiet))) = | |
| 34128 
8650a073dd9b
made Quickcheck take structured proof assumptions into account (like Refute and Nitpick) by default;
 blanchet parents: 
33583diff
changeset | 71 |   Test_Params { size = size, iterations = iterations, default_type = default_type,
 | 
| 35379 | 72 | no_assms = no_assms, report = report, quiet = quiet }; | 
| 73 | fun map_test_params f (Test_Params { size, iterations, default_type, no_assms, report, quiet }) =
 | |
| 74 | make_test_params (f ((size, iterations), ((default_type, no_assms), (report, quiet)))); | |
| 34128 
8650a073dd9b
made Quickcheck take structured proof assumptions into account (like Refute and Nitpick) by default;
 blanchet parents: 
33583diff
changeset | 75 | fun merge_test_params (Test_Params { size = size1, iterations = iterations1, default_type = default_type1,
 | 
| 35379 | 76 | no_assms = no_assms1, report = report1, quiet = quiet1 }, | 
| 34128 
8650a073dd9b
made Quickcheck take structured proof assumptions into account (like Refute and Nitpick) by default;
 blanchet parents: 
33583diff
changeset | 77 |   Test_Params { size = size2, iterations = iterations2, default_type = default_type2,
 | 
| 35379 | 78 | no_assms = no_assms2, report = report2, quiet = quiet2 }) = | 
| 31599 | 79 | make_test_params ((Int.max (size1, size2), Int.max (iterations1, iterations2)), | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 80 | ((case default_type1 of NONE => default_type2 | _ => default_type1, no_assms1 orelse no_assms2), | 
| 35379 | 81 | (report1 orelse report2, quiet1 orelse quiet2))); | 
| 28309 | 82 | |
| 33522 | 83 | structure Data = Theory_Data | 
| 84 | ( | |
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 85 | type T = (string * (Proof.context -> bool -> term -> int -> term list option * (bool list * bool))) list | 
| 28309 | 86 | * test_params; | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 87 | val empty = ([], Test_Params | 
| 35379 | 88 |     { size = 10, iterations = 100, default_type = NONE, no_assms = false, report = false, quiet = false});
 | 
| 28256 | 89 | val extend = I; | 
| 33522 | 90 | fun merge ((generators1, params1), (generators2, params2)) : T = | 
| 91 | (AList.merge (op =) (K true) (generators1, generators2), | |
| 28309 | 92 | merge_test_params (params1, params2)); | 
| 33522 | 93 | ); | 
| 28256 | 94 | |
| 28309 | 95 | val add_generator = Data.map o apfst o AList.update (op =); | 
| 96 | ||
| 28315 | 97 | (* generating tests *) | 
| 98 | ||
| 28309 | 99 | fun mk_tester_select name ctxt = | 
| 100 | case AList.lookup (op =) ((fst o Data.get o ProofContext.theory_of) ctxt) name | |
| 101 |    of NONE => error ("No such quickcheck generator: " ^ name)
 | |
| 102 | | SOME generator => generator ctxt; | |
| 103 | ||
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 104 | fun mk_testers ctxt report t = | 
| 28309 | 105 | (map snd o fst o Data.get o ProofContext.theory_of) ctxt | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 106 | |> map_filter (fn generator => try (generator ctxt report) t); | 
| 28309 | 107 | |
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 108 | fun mk_testers_strict ctxt report t = | 
| 28309 | 109 | let | 
| 110 | val generators = ((map snd o fst o Data.get o ProofContext.theory_of) ctxt) | |
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 111 | val testers = map (fn generator => Exn.capture (generator ctxt report) t) generators; | 
| 28309 | 112 | in if forall (is_none o Exn.get_result) testers | 
| 113 | then [(Exn.release o snd o split_last) testers] | |
| 114 | else map_filter Exn.get_result testers | |
| 115 | end; | |
| 116 | ||
| 28315 | 117 | |
| 118 | (* testing propositions *) | |
| 119 | ||
| 28309 | 120 | fun prep_test_term t = | 
| 121 | let | |
| 29266 | 122 | val _ = (null (Term.add_tvars t []) andalso null (Term.add_tfrees t [])) orelse | 
| 28309 | 123 | error "Term to be tested contains type variables"; | 
| 29266 | 124 | val _ = null (Term.add_vars t []) orelse | 
| 28309 | 125 | error "Term to be tested contains schematic variables"; | 
| 31138 | 126 | val frees = Term.add_frees t []; | 
| 28309 | 127 | in (map fst frees, list_abs_free (frees, t)) end | 
| 28256 | 128 | |
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35077diff
changeset | 129 | fun cpu_time description f = | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35077diff
changeset | 130 | let | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35077diff
changeset | 131 | val start = start_timing () | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35077diff
changeset | 132 | val result = Exn.capture f () | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35077diff
changeset | 133 | val time = Time.toMilliseconds (#cpu (end_timing start)) | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35077diff
changeset | 134 | in (Exn.release result, (description, time)) end | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35077diff
changeset | 135 | |
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 136 | fun gen_test_term ctxt quiet report generator_name size i t = | 
| 28309 | 137 | let | 
| 138 | val (names, t') = prep_test_term t; | |
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35077diff
changeset | 139 | val (testers, comp_time) = cpu_time "quickcheck compilation" | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35077diff
changeset | 140 | (fn () => (case generator_name | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 141 | of NONE => if quiet then mk_testers ctxt report t' else mk_testers_strict ctxt report t' | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 142 | | SOME name => [mk_tester_select name ctxt report t'])); | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 143 | fun iterate f 0 report = (NONE, report) | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 144 | | iterate f j report = | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 145 | let | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 146 | val (test_result, single_report) = apsnd Run (f ()) handle Match => (if quiet then () | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 147 | else warning "Exception Match raised during quickcheck"; (NONE, MatchExc)) | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 148 | val report = collect_single_report single_report report | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 149 | in | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 150 | case test_result of NONE => iterate f (j - 1) report | SOME q => (SOME q, report) | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 151 | end | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 152 |     val empty_report = Report { iterations = 0, raised_match_errors = 0,
 | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 153 | satisfied_assms = [], positive_concl_tests = 0 } | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 154 | fun with_testers k [] = (NONE, []) | 
| 28309 | 155 | | with_testers k (tester :: testers) = | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 156 | case iterate (fn () => tester (k - 1)) i empty_report | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 157 | of (NONE, report) => apsnd (cons report) (with_testers k testers) | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 158 | | (SOME q, report) => (SOME q, [report]); | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 159 | fun with_size k reports = if k > size then (NONE, reports) | 
| 28309 | 160 |       else (if quiet then () else priority ("Test data size: " ^ string_of_int k);
 | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 161 | let | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 162 | val (result, new_report) = with_testers k testers | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 163 | val reports = ((k, new_report) :: reports) | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 164 | in case result of NONE => with_size (k + 1) reports | SOME q => (SOME q, reports) end); | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 165 | val ((result, reports), exec_time) = cpu_time "quickcheck execution" | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 166 | (fn () => apfst | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 167 | (fn result => case result of NONE => NONE | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 168 | | SOME ts => SOME (names ~~ ts)) (with_size 1 [])) | 
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34128diff
changeset | 169 | in | 
| 35380 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35379diff
changeset | 170 | (result, ([exec_time, comp_time], if report then SOME reports else NONE)) | 
| 28309 | 171 | end; | 
| 172 | ||
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35077diff
changeset | 173 | fun test_term ctxt quiet generator_name size i t = | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 174 | fst (gen_test_term ctxt quiet false generator_name size i t) | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35077diff
changeset | 175 | |
| 28309 | 176 | fun monomorphic_term thy insts default_T = | 
| 177 | let | |
| 178 | fun subst (T as TFree (v, S)) = | |
| 179 | let | |
| 180 | val T' = AList.lookup (op =) insts v | |
| 181 | |> the_default (the_default T default_T) | |
| 28315 | 182 | in if Sign.of_sort thy (T, S) then T' | 
| 28309 | 183 |             else error ("Type " ^ Syntax.string_of_typ_global thy T ^
 | 
| 184 | " to be substituted for variable " ^ | |
| 185 | Syntax.string_of_typ_global thy T ^ "\ndoes not have sort " ^ | |
| 186 | Syntax.string_of_sort_global thy S) | |
| 187 | end | |
| 188 | | subst T = T; | |
| 189 | in (map_types o map_atyps) subst end; | |
| 190 | ||
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 191 | fun test_goal quiet report generator_name size iterations default_T no_assms insts i assms state = | 
| 28309 | 192 | let | 
| 193 | val ctxt = Proof.context_of state; | |
| 194 | val thy = Proof.theory_of state; | |
| 195 |     fun strip (Const ("all", _) $ Abs (_, _, t)) = strip t
 | |
| 196 | | strip t = t; | |
| 33291 | 197 |     val {goal = st, ...} = Proof.raw_goal state;
 | 
| 28309 | 198 | val (gi, frees) = Logic.goal_params (prop_of st) i; | 
| 34128 
8650a073dd9b
made Quickcheck take structured proof assumptions into account (like Refute and Nitpick) by default;
 blanchet parents: 
33583diff
changeset | 199 | val gi' = Logic.list_implies (if no_assms then [] else assms, | 
| 
8650a073dd9b
made Quickcheck take structured proof assumptions into account (like Refute and Nitpick) by default;
 blanchet parents: 
33583diff
changeset | 200 | subst_bounds (frees, strip gi)) | 
| 28309 | 201 | |> monomorphic_term thy insts default_T | 
| 35625 | 202 | |> Object_Logic.atomize_term thy; | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 203 | in gen_test_term ctxt quiet report generator_name size iterations gi' end; | 
| 28315 | 204 | |
| 35077 
c1dac8ace020
make Quickcheck identify itself, so people don't submit bug reports to me thinking that it was Nitpick
 blanchet parents: 
34948diff
changeset | 205 | fun pretty_counterex ctxt NONE = Pretty.str "Quickcheck found no counterexample." | 
| 28315 | 206 | | pretty_counterex ctxt (SOME cex) = | 
| 35077 
c1dac8ace020
make Quickcheck identify itself, so people don't submit bug reports to me thinking that it was Nitpick
 blanchet parents: 
34948diff
changeset | 207 | Pretty.chunks (Pretty.str "Quickcheck found a counterexample:\n" :: | 
| 28315 | 208 | map (fn (s, t) => | 
| 209 | Pretty.block [Pretty.str (s ^ " ="), Pretty.brk 1, Syntax.pretty_term ctxt t]) cex); | |
| 210 | ||
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 211 | fun pretty_report (Report {iterations = iterations, raised_match_errors = raised_match_errors,
 | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 212 | satisfied_assms = satisfied_assms, positive_concl_tests = positive_concl_tests}) = | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 213 | let | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 214 | fun pretty_stat s i = Pretty.block ([Pretty.str (s ^ ": " ^ string_of_int i)]) | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 215 | in | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 216 | ([pretty_stat "iterations" iterations, | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 217 | pretty_stat "match exceptions" raised_match_errors] | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 218 |      @ map_index (fn (i, n) => pretty_stat ("satisfied " ^ string_of_int (i + 1) ^ ". assumption") n)
 | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 219 | satisfied_assms | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 220 | @ [pretty_stat "positive conclusion tests" positive_concl_tests]) | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 221 | end | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 222 | |
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 223 | fun pretty_reports' [report] = [Pretty.chunks (pretty_report report)] | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 224 | | pretty_reports' reports = | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 225 | map_index (fn (i, report) => | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 226 | Pretty.chunks (Pretty.str (string_of_int (i + 1) ^ ". generator:\n") :: pretty_report report)) | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 227 | reports | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 228 | |
| 35380 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35379diff
changeset | 229 | fun pretty_reports ctxt (SOME reports) = | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 230 | Pretty.chunks (Pretty.str "Quickcheck report:" :: | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 231 | maps (fn (size, reports) => | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 232 |       Pretty.str ("size " ^ string_of_int size ^ ":") :: pretty_reports' reports @ [Pretty.brk 1])
 | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 233 | (rev reports)) | 
| 35380 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35379diff
changeset | 234 | | pretty_reports ctxt NONE = Pretty.str "" | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 235 | |
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 236 | fun pretty_counterex_and_reports ctxt (cex, (timing, reports)) = | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 237 | Pretty.chunks [pretty_counterex ctxt cex, pretty_reports ctxt reports] | 
| 28315 | 238 | |
| 239 | (* automatic testing *) | |
| 28309 | 240 | |
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 241 | fun auto_quickcheck state = | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 242 | if not (!auto) then | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 243 | (false, state) | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 244 | else | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 245 | let | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 246 | val ctxt = Proof.context_of state; | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 247 | val assms = map term_of (Assumption.all_assms_of ctxt); | 
| 35379 | 248 |       val Test_Params { size, iterations, default_type, no_assms, report, quiet } =
 | 
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 249 | (snd o Data.get o Proof.theory_of) state; | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 250 | val res = | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 251 | try (test_goal true false NONE size iterations default_type no_assms [] 1 assms) state; | 
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 252 | in | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 253 | case res of | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 254 | NONE => (false, state) | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 255 | | SOME (NONE, report) => (false, state) | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 256 | | SOME (cex, report) => (true, Proof.goal_message (K (Pretty.chunks [Pretty.str "", | 
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 257 | Pretty.mark Markup.hilite (pretty_counterex ctxt cex)])) state) | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 258 | end | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 259 | |
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 260 | val setup = Auto_Counterexample.register_tool ("quickcheck", auto_quickcheck)
 | 
| 28315 | 261 | |
| 262 | ||
| 30980 | 263 | (* Isar commands *) | 
| 28315 | 264 | |
| 28336 | 265 | fun read_nat s = case (Library.read_int o Symbol.explode) s | 
| 266 | of (k, []) => if k >= 0 then k | |
| 267 |       else error ("Not a natural number: " ^ s)
 | |
| 268 |   | (_, _ :: _) => error ("Not a natural number: " ^ s);
 | |
| 34128 
8650a073dd9b
made Quickcheck take structured proof assumptions into account (like Refute and Nitpick) by default;
 blanchet parents: 
33583diff
changeset | 269 | fun read_bool "false" = false | 
| 
8650a073dd9b
made Quickcheck take structured proof assumptions into account (like Refute and Nitpick) by default;
 blanchet parents: 
33583diff
changeset | 270 | | read_bool "true" = true | 
| 
8650a073dd9b
made Quickcheck take structured proof assumptions into account (like Refute and Nitpick) by default;
 blanchet parents: 
33583diff
changeset | 271 |   | read_bool s = error ("Not a Boolean value: " ^ s)
 | 
| 28315 | 272 | |
| 28336 | 273 | fun parse_test_param ctxt ("size", arg) =
 | 
| 274 | (apfst o apfst o K) (read_nat arg) | |
| 275 |   | parse_test_param ctxt ("iterations", arg) =
 | |
| 276 | (apfst o apsnd o K) (read_nat arg) | |
| 277 |   | parse_test_param ctxt ("default_type", arg) =
 | |
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 278 | (apsnd o apfst o apfst o K o SOME) (ProofContext.read_typ ctxt arg) | 
| 34128 
8650a073dd9b
made Quickcheck take structured proof assumptions into account (like Refute and Nitpick) by default;
 blanchet parents: 
33583diff
changeset | 279 |   | parse_test_param ctxt ("no_assms", arg) =
 | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 280 | (apsnd o apfst o apsnd o K) (read_bool arg) | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 281 |   | parse_test_param ctxt ("report", arg) =
 | 
| 35379 | 282 | (apsnd o apsnd o apfst o K) (read_bool arg) | 
| 283 |   | parse_test_param ctxt ("quiet", arg) =
 | |
| 284 | (apsnd o apsnd o apsnd o K) (read_bool arg) | |
| 28336 | 285 | | parse_test_param ctxt (name, _) = | 
| 34128 
8650a073dd9b
made Quickcheck take structured proof assumptions into account (like Refute and Nitpick) by default;
 blanchet parents: 
33583diff
changeset | 286 |       error ("Unknown test parameter: " ^ name);
 | 
| 28315 | 287 | |
| 28336 | 288 | fun parse_test_param_inst ctxt ("generator", arg) =
 | 
| 289 | (apsnd o apfst o K o SOME) arg | |
| 290 | | parse_test_param_inst ctxt (name, arg) = | |
| 291 | case try (ProofContext.read_typ ctxt) name | |
| 292 | of SOME (TFree (v, _)) => (apsnd o apsnd o AList.update (op =)) | |
| 293 | (v, ProofContext.read_typ ctxt arg) | |
| 294 | | _ => (apfst o parse_test_param ctxt) (name, arg); | |
| 28309 | 295 | |
| 28336 | 296 | fun quickcheck_params_cmd args thy = | 
| 28315 | 297 | let | 
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
35625diff
changeset | 298 | val ctxt = ProofContext.init_global thy; | 
| 28336 | 299 | val f = fold (parse_test_param ctxt) args; | 
| 28315 | 300 | in | 
| 301 | thy | |
| 28336 | 302 | |> (Data.map o apsnd o map_test_params) f | 
| 28315 | 303 | end; | 
| 304 | ||
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 305 | fun gen_quickcheck args i state = | 
| 28315 | 306 | let | 
| 32297 | 307 | val thy = Proof.theory_of state; | 
| 308 | val ctxt = Proof.context_of state; | |
| 34128 
8650a073dd9b
made Quickcheck take structured proof assumptions into account (like Refute and Nitpick) by default;
 blanchet parents: 
33583diff
changeset | 309 | val assms = map term_of (Assumption.all_assms_of ctxt); | 
| 28315 | 310 | val default_params = (dest_test_params o snd o Data.get) thy; | 
| 28336 | 311 | val f = fold (parse_test_param_inst ctxt) args; | 
| 35379 | 312 | val (((size, iterations), ((default_type, no_assms), (report, quiet))), (generator_name, insts)) = | 
| 28336 | 313 | f (default_params, (NONE, [])); | 
| 32297 | 314 | in | 
| 35379 | 315 | test_goal quiet report generator_name size iterations default_type no_assms insts i assms state | 
| 32297 | 316 | end; | 
| 317 | ||
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 318 | fun quickcheck args i state = fst (gen_quickcheck args i state); | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 319 | |
| 32297 | 320 | fun quickcheck_cmd args i state = | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 321 | gen_quickcheck args i (Toplevel.proof_of state) | 
| 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 322 | |> Pretty.writeln o pretty_counterex_and_reports (Toplevel.context_of state); | 
| 28309 | 323 | |
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 324 | val parse_arg = Parse.name -- (Scan.optional (Parse.$$$ "=" |-- Parse.name) "true"); | 
| 28309 | 325 | |
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 326 | val parse_args = Parse.$$$ "[" |-- Parse.list1 parse_arg --| Parse.$$$ "]" | 
| 28336 | 327 | || Scan.succeed []; | 
| 328 | ||
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 329 | val _ = | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 330 | Outer_Syntax.command "quickcheck_params" "set parameters for random testing" Keyword.thy_decl | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 331 | (parse_args >> (fn args => Toplevel.theory (quickcheck_params_cmd args))); | 
| 28309 | 332 | |
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 333 | val _ = | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 334 | Outer_Syntax.improper_command "quickcheck" "try to find counterexample for subgoal" Keyword.diag | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 335 | (parse_args -- Scan.optional Parse.nat 1 | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 336 | >> (fn (args, i) => Toplevel.no_timing o Toplevel.keep (quickcheck_cmd args i))); | 
| 28309 | 337 | |
| 28315 | 338 | end; | 
| 28309 | 339 | |
| 340 | ||
| 28315 | 341 | val auto_quickcheck = Quickcheck.auto; |