| author | wenzelm | 
| Fri, 20 Nov 2009 15:33:10 +0100 | |
| changeset 33820 | 082d9bc6992d | 
| parent 33583 | b5e0909cd5ea | 
| child 34128 | 8650a073dd9b | 
| 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 | 
| 30980 | 10 | val test_term: Proof.context -> bool -> string option -> int -> int -> term -> | 
| 11 | (string * term) list option | |
| 12 | val add_generator: string * (Proof.context -> term -> int -> term list option) -> theory -> theory | |
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 13 | val setup: theory -> theory | 
| 32297 | 14 | val quickcheck: (string * string) list -> int -> Proof.state -> (string * term) list option | 
| 28256 | 15 | end; | 
| 16 | ||
| 17 | structure Quickcheck : QUICKCHECK = | |
| 18 | struct | |
| 19 | ||
| 30980 | 20 | (* preferences *) | 
| 21 | ||
| 32740 | 22 | val auto = Unsynchronized.ref false; | 
| 30980 | 23 | |
| 24 | val _ = | |
| 25 | ProofGeneralPgip.add_preference Preferences.category_tracing | |
| 32966 
5b21661fe618
indicate CRITICAL nature of various setmp combinators;
 wenzelm parents: 
32859diff
changeset | 26 | (setmp_CRITICAL auto true (fn () => | 
| 30980 | 27 | Preferences.bool_pref auto | 
| 28 | "auto-quickcheck" | |
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 29 | "Whether to run Quickcheck automatically.") ()); | 
| 30980 | 30 | |
| 30973 
304ab57afa6e
observe distinction between Pure/Tools and Tools more closely
 haftmann parents: 
30824diff
changeset | 31 | |
| 28315 | 32 | (* quickcheck configuration -- default parameters, test generators *) | 
| 33 | ||
| 28309 | 34 | datatype test_params = Test_Params of | 
| 35 |   { size: int, iterations: int, default_type: typ option };
 | |
| 36 | ||
| 30973 
304ab57afa6e
observe distinction between Pure/Tools and Tools more closely
 haftmann parents: 
30824diff
changeset | 37 | fun dest_test_params (Test_Params { size, iterations, default_type }) =
 | 
| 28315 | 38 | ((size, iterations), default_type); | 
| 31599 | 39 | fun make_test_params ((size, iterations), default_type) = | 
| 28309 | 40 |   Test_Params { size = size, iterations = iterations, default_type = default_type };
 | 
| 41 | fun map_test_params f (Test_Params { size, iterations, default_type}) =
 | |
| 31599 | 42 | make_test_params (f ((size, iterations), default_type)); | 
| 30973 
304ab57afa6e
observe distinction between Pure/Tools and Tools more closely
 haftmann parents: 
30824diff
changeset | 43 | fun merge_test_params (Test_Params { size = size1, iterations = iterations1, default_type = default_type1 },
 | 
| 
304ab57afa6e
observe distinction between Pure/Tools and Tools more closely
 haftmann parents: 
30824diff
changeset | 44 |   Test_Params { size = size2, iterations = iterations2, default_type = default_type2 }) =
 | 
| 31599 | 45 | make_test_params ((Int.max (size1, size2), Int.max (iterations1, iterations2)), | 
| 28309 | 46 | case default_type1 of NONE => default_type2 | _ => default_type1); | 
| 47 | ||
| 33522 | 48 | structure Data = Theory_Data | 
| 49 | ( | |
| 28309 | 50 | type T = (string * (Proof.context -> term -> int -> term list option)) list | 
| 51 | * test_params; | |
| 52 |   val empty = ([], Test_Params { size = 10, iterations = 100, default_type = NONE });
 | |
| 28256 | 53 | val extend = I; | 
| 33522 | 54 | fun merge ((generators1, params1), (generators2, params2)) : T = | 
| 55 | (AList.merge (op =) (K true) (generators1, generators2), | |
| 28309 | 56 | merge_test_params (params1, params2)); | 
| 33522 | 57 | ); | 
| 28256 | 58 | |
| 28309 | 59 | val add_generator = Data.map o apfst o AList.update (op =); | 
| 60 | ||
| 28315 | 61 | |
| 62 | (* generating tests *) | |
| 63 | ||
| 28309 | 64 | fun mk_tester_select name ctxt = | 
| 65 | case AList.lookup (op =) ((fst o Data.get o ProofContext.theory_of) ctxt) name | |
| 66 |    of NONE => error ("No such quickcheck generator: " ^ name)
 | |
| 67 | | SOME generator => generator ctxt; | |
| 68 | ||
| 69 | fun mk_testers ctxt t = | |
| 70 | (map snd o fst o Data.get o ProofContext.theory_of) ctxt | |
| 71 | |> map_filter (fn generator => try (generator ctxt) t); | |
| 72 | ||
| 73 | fun mk_testers_strict ctxt t = | |
| 74 | let | |
| 75 | val generators = ((map snd o fst o Data.get o ProofContext.theory_of) ctxt) | |
| 76 | val testers = map (fn generator => Exn.capture (generator ctxt) t) generators; | |
| 77 | in if forall (is_none o Exn.get_result) testers | |
| 78 | then [(Exn.release o snd o split_last) testers] | |
| 79 | else map_filter Exn.get_result testers | |
| 80 | end; | |
| 81 | ||
| 28315 | 82 | |
| 83 | (* testing propositions *) | |
| 84 | ||
| 28309 | 85 | fun prep_test_term t = | 
| 86 | let | |
| 29266 | 87 | val _ = (null (Term.add_tvars t []) andalso null (Term.add_tfrees t [])) orelse | 
| 28309 | 88 | error "Term to be tested contains type variables"; | 
| 29266 | 89 | val _ = null (Term.add_vars t []) orelse | 
| 28309 | 90 | error "Term to be tested contains schematic variables"; | 
| 31138 | 91 | val frees = Term.add_frees t []; | 
| 28309 | 92 | in (map fst frees, list_abs_free (frees, t)) end | 
| 28256 | 93 | |
| 28315 | 94 | fun test_term ctxt quiet generator_name size i t = | 
| 28309 | 95 | let | 
| 96 | val (names, t') = prep_test_term t; | |
| 97 | val testers = case generator_name | |
| 98 | of NONE => if quiet then mk_testers ctxt t' else mk_testers_strict ctxt t' | |
| 99 | | SOME name => [mk_tester_select name ctxt t']; | |
| 100 | fun iterate f 0 = NONE | |
| 31153 | 101 | | iterate f j = case f () handle Match => (if quiet then () | 
| 28309 | 102 | else warning "Exception Match raised during quickcheck"; NONE) | 
| 31153 | 103 | of NONE => iterate f (j - 1) | SOME q => SOME q; | 
| 28309 | 104 | fun with_testers k [] = NONE | 
| 105 | | with_testers k (tester :: testers) = | |
| 31153 | 106 | case iterate (fn () => tester (k - 1)) i | 
| 28309 | 107 | of NONE => with_testers k testers | 
| 108 | | SOME q => SOME q; | |
| 109 | fun with_size k = if k > size then NONE | |
| 110 |       else (if quiet then () else priority ("Test data size: " ^ string_of_int k);
 | |
| 111 | case with_testers k testers | |
| 112 | of NONE => with_size (k + 1) | SOME q => SOME q); | |
| 113 | in case with_size 1 | |
| 114 | of NONE => NONE | |
| 115 | | SOME ts => SOME (names ~~ ts) | |
| 116 | end; | |
| 117 | ||
| 118 | fun monomorphic_term thy insts default_T = | |
| 119 | let | |
| 120 | fun subst (T as TFree (v, S)) = | |
| 121 | let | |
| 122 | val T' = AList.lookup (op =) insts v | |
| 123 | |> the_default (the_default T default_T) | |
| 28315 | 124 | in if Sign.of_sort thy (T, S) then T' | 
| 28309 | 125 |             else error ("Type " ^ Syntax.string_of_typ_global thy T ^
 | 
| 126 | " to be substituted for variable " ^ | |
| 127 | Syntax.string_of_typ_global thy T ^ "\ndoes not have sort " ^ | |
| 128 | Syntax.string_of_sort_global thy S) | |
| 129 | end | |
| 130 | | subst T = T; | |
| 131 | in (map_types o map_atyps) subst end; | |
| 132 | ||
| 28315 | 133 | fun test_goal quiet generator_name size iterations default_T insts i assms state = | 
| 28309 | 134 | let | 
| 135 | val ctxt = Proof.context_of state; | |
| 136 | val thy = Proof.theory_of state; | |
| 137 |     fun strip (Const ("all", _) $ Abs (_, _, t)) = strip t
 | |
| 138 | | strip t = t; | |
| 33291 | 139 |     val {goal = st, ...} = Proof.raw_goal state;
 | 
| 28309 | 140 | val (gi, frees) = Logic.goal_params (prop_of st) i; | 
| 141 | val gi' = Logic.list_implies (assms, subst_bounds (frees, strip gi)) | |
| 142 | |> monomorphic_term thy insts default_T | |
| 143 | |> ObjectLogic.atomize_term thy; | |
| 28315 | 144 | in test_term ctxt quiet generator_name size iterations gi' end; | 
| 145 | ||
| 146 | fun pretty_counterex ctxt NONE = Pretty.str "No counterexamples found." | |
| 147 | | pretty_counterex ctxt (SOME cex) = | |
| 148 | Pretty.chunks (Pretty.str "Counterexample found:\n" :: | |
| 149 | map (fn (s, t) => | |
| 150 | Pretty.block [Pretty.str (s ^ " ="), Pretty.brk 1, Syntax.pretty_term ctxt t]) cex); | |
| 151 | ||
| 152 | ||
| 153 | (* automatic testing *) | |
| 28309 | 154 | |
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 155 | fun auto_quickcheck state = | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 156 | if not (!auto) then | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 157 | (false, state) | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 158 | else | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 159 | let | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 160 | val ctxt = Proof.context_of state; | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 161 | val assms = map term_of (Assumption.all_assms_of ctxt); | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 162 |       val Test_Params { size, iterations, default_type } =
 | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 163 | (snd o Data.get o Proof.theory_of) state; | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 164 | val res = | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 165 | try (test_goal true NONE size iterations default_type [] 1 assms) state; | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 166 | in | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 167 | case res of | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 168 | NONE => (false, state) | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 169 | | SOME NONE => (false, state) | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 170 | | SOME cex => (true, Proof.goal_message (K (Pretty.chunks [Pretty.str "", | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 171 | Pretty.mark Markup.hilite (pretty_counterex ctxt cex)])) state) | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 172 | end | 
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 173 | |
| 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
33560diff
changeset | 174 | val setup = Auto_Counterexample.register_tool ("quickcheck", auto_quickcheck)
 | 
| 28315 | 175 | |
| 176 | ||
| 30980 | 177 | (* Isar commands *) | 
| 28315 | 178 | |
| 28336 | 179 | fun read_nat s = case (Library.read_int o Symbol.explode) s | 
| 180 | of (k, []) => if k >= 0 then k | |
| 181 |       else error ("Not a natural number: " ^ s)
 | |
| 182 |   | (_, _ :: _) => error ("Not a natural number: " ^ s);
 | |
| 28315 | 183 | |
| 28336 | 184 | fun parse_test_param ctxt ("size", arg) =
 | 
| 185 | (apfst o apfst o K) (read_nat arg) | |
| 186 |   | parse_test_param ctxt ("iterations", arg) =
 | |
| 187 | (apfst o apsnd o K) (read_nat arg) | |
| 188 |   | parse_test_param ctxt ("default_type", arg) =
 | |
| 189 | (apsnd o K o SOME) (ProofContext.read_typ ctxt arg) | |
| 190 | | parse_test_param ctxt (name, _) = | |
| 191 |       error ("Bad test parameter: " ^ name);
 | |
| 28315 | 192 | |
| 28336 | 193 | fun parse_test_param_inst ctxt ("generator", arg) =
 | 
| 194 | (apsnd o apfst o K o SOME) arg | |
| 195 | | parse_test_param_inst ctxt (name, arg) = | |
| 196 | case try (ProofContext.read_typ ctxt) name | |
| 197 | of SOME (TFree (v, _)) => (apsnd o apsnd o AList.update (op =)) | |
| 198 | (v, ProofContext.read_typ ctxt arg) | |
| 199 | | _ => (apfst o parse_test_param ctxt) (name, arg); | |
| 28309 | 200 | |
| 28336 | 201 | fun quickcheck_params_cmd args thy = | 
| 28315 | 202 | let | 
| 203 | val ctxt = ProofContext.init thy; | |
| 28336 | 204 | val f = fold (parse_test_param ctxt) args; | 
| 28315 | 205 | in | 
| 206 | thy | |
| 28336 | 207 | |> (Data.map o apsnd o map_test_params) f | 
| 28315 | 208 | end; | 
| 209 | ||
| 32297 | 210 | fun quickcheck args i state = | 
| 28315 | 211 | let | 
| 32297 | 212 | val thy = Proof.theory_of state; | 
| 213 | val ctxt = Proof.context_of state; | |
| 28315 | 214 | val default_params = (dest_test_params o snd o Data.get) thy; | 
| 28336 | 215 | val f = fold (parse_test_param_inst ctxt) args; | 
| 28315 | 216 | val (((size, iterations), default_type), (generator_name, insts)) = | 
| 28336 | 217 | f (default_params, (NONE, [])); | 
| 32297 | 218 | in | 
| 219 | test_goal false generator_name size iterations default_type insts i [] state | |
| 220 | end; | |
| 221 | ||
| 222 | fun quickcheck_cmd args i state = | |
| 223 | quickcheck args i (Toplevel.proof_of state) | |
| 224 | |> Pretty.writeln o pretty_counterex (Toplevel.context_of state); | |
| 28309 | 225 | |
| 226 | local structure P = OuterParse and K = OuterKeyword in | |
| 227 | ||
| 28336 | 228 | val parse_arg = P.name --| P.$$$ "=" -- P.name; | 
| 229 | val parse_args = P.$$$ "[" |-- P.list1 parse_arg --| P.$$$ "]" | |
| 230 | || Scan.succeed []; | |
| 231 | ||
| 28315 | 232 | val _ = OuterSyntax.command "quickcheck_params" "set parameters for random testing" K.thy_decl | 
| 28336 | 233 | (parse_args >> (fn args => Toplevel.theory (quickcheck_params_cmd args))); | 
| 28309 | 234 | |
| 28315 | 235 | val _ = OuterSyntax.improper_command "quickcheck" "try to find counterexample for subgoal" K.diag | 
| 28336 | 236 | (parse_args -- Scan.optional P.nat 1 | 
| 237 | >> (fn (args, i) => Toplevel.no_timing o Toplevel.keep (quickcheck_cmd args i))); | |
| 28309 | 238 | |
| 239 | end; (*local*) | |
| 240 | ||
| 28315 | 241 | end; | 
| 28309 | 242 | |
| 243 | ||
| 28315 | 244 | val auto_quickcheck = Quickcheck.auto; |