src/Tools/quickcheck.ML
author bulwahn
Fri, 11 Feb 2011 11:47:42 +0100
changeset 42624 dbd00d8a4784
parent 42606 bd7ee90267f2
child 42625 aa94a003dcdf
permissions -rw-r--r--
quickcheck invokes TimeLimit.timeLimit only in one separate function
     1 (*  Title:      Tools/quickcheck.ML
     2     Author:     Stefan Berghofer, Florian Haftmann, Lukas Bulwahn, TU Muenchen
     3 
     4 Generic counterexample search engine.
     5 *)
     6 
     7 signature QUICKCHECK =
     8 sig
     9   val setup: theory -> theory
    10   (* configuration *)
    11   val auto: bool Unsynchronized.ref
    12   val timing : bool Unsynchronized.ref
    13   val tester : string Config.T
    14   val size : int Config.T
    15   val iterations : int Config.T
    16   val no_assms : bool Config.T
    17   val report : bool Config.T
    18   val quiet : bool Config.T
    19   val timeout : real Config.T
    20   val finite_types : bool Config.T
    21   val finite_type_size : int Config.T
    22   datatype report = Report of
    23     { iterations : int, raised_match_errors : int,
    24       satisfied_assms : int list, positive_concl_tests : int }
    25   datatype expectation = No_Expectation | No_Counterexample | Counterexample;
    26   datatype test_params = Test_Params of {default_type: typ list, expect : expectation};
    27   val test_params_of : Proof.context -> test_params
    28   val map_test_params : (typ list * expectation -> typ list * expectation)
    29     -> Context.generic -> Context.generic
    30   val add_generator:
    31     string * (Proof.context -> term -> int -> term list option * report option)
    32       -> Context.generic -> Context.generic
    33   (* testing terms and proof states *)
    34   val test_term: Proof.context -> bool -> term ->
    35     (string * term) list option * ((string * int) list * ((int * report) list) option)
    36   val test_goal_terms:
    37     Proof.context -> bool -> (string * typ) list -> term list
    38       -> (string * term) list option * ((string * int) list * ((int * report) list) option) list
    39   val quickcheck: (string * string list) list -> int -> Proof.state -> (string * term) list option
    40 end;
    41 
    42 structure Quickcheck : QUICKCHECK =
    43 struct
    44 
    45 (* preferences *)
    46 
    47 val auto = Unsynchronized.ref false;
    48 
    49 val timing = Unsynchronized.ref false;
    50 
    51 val _ =
    52   ProofGeneralPgip.add_preference Preferences.category_tracing
    53   (Unsynchronized.setmp auto true (fn () =>
    54     Preferences.bool_pref auto
    55       "auto-quickcheck"
    56       "Run Quickcheck automatically.") ());
    57 
    58 (* quickcheck report *)
    59 
    60 datatype report = Report of
    61   { iterations : int, raised_match_errors : int,
    62     satisfied_assms : int list, positive_concl_tests : int }
    63 
    64 (* expectation *)
    65 
    66 datatype expectation = No_Expectation | No_Counterexample | Counterexample;
    67 
    68 fun merge_expectation (expect1, expect2) =
    69   if expect1 = expect2 then expect1 else No_Expectation
    70 
    71 (* quickcheck configuration -- default parameters, test generators *)
    72 val (tester, setup_tester) = Attrib.config_string "quickcheck_tester" (K "")
    73 val (size, setup_size) = Attrib.config_int "quickcheck_size" (K 10)
    74 val (iterations, setup_iterations) = Attrib.config_int "quickcheck_iterations" (K 100)
    75 val (no_assms, setup_no_assms) = Attrib.config_bool "quickcheck_no_assms" (K false)
    76 val (report, setup_report) = Attrib.config_bool "quickcheck_report" (K true)
    77 val (quiet, setup_quiet) = Attrib.config_bool "quickcheck_quiet" (K false)
    78 val (timeout, setup_timeout) = Attrib.config_real "quickcheck_timeout" (K 30.0)
    79 val (finite_types, setup_finite_types) = Attrib.config_bool "quickcheck_finite_types" (K true)
    80 val (finite_type_size, setup_finite_type_size) =
    81   Attrib.config_int "quickcheck_finite_type_size" (K 3)
    82 
    83 val setup_config =
    84   setup_tester #> setup_size #> setup_iterations #> setup_no_assms #> setup_report #> setup_quiet
    85     #> setup_timeout #> setup_finite_types #> setup_finite_type_size
    86 
    87 datatype test_params = Test_Params of
    88   {default_type: typ list, expect : expectation};
    89 
    90 fun dest_test_params (Test_Params {default_type, expect}) = (default_type, expect);
    91 
    92 fun make_test_params (default_type, expect) =
    93   Test_Params {default_type = default_type, expect = expect};
    94 
    95 fun map_test_params' f (Test_Params {default_type, expect}) =
    96   make_test_params (f (default_type, expect));
    97 
    98 fun merge_test_params
    99   (Test_Params {default_type = default_type1, expect = expect1},
   100     Test_Params {default_type = default_type2, expect = expect2}) =
   101   make_test_params
   102     (merge (op =) (default_type1, default_type2), merge_expectation (expect1, expect2));
   103 
   104 structure Data = Generic_Data
   105 (
   106   type T =
   107     (string * (Proof.context -> term -> int -> term list option * report option)) list
   108       * test_params;
   109   val empty = ([], Test_Params {default_type = [], expect = No_Expectation});
   110   val extend = I;
   111   fun merge ((generators1, params1), (generators2, params2)) : T =
   112     (AList.merge (op =) (K true) (generators1, generators2),
   113       merge_test_params (params1, params2));
   114 );
   115 
   116 val test_params_of = snd o Data.get o Context.Proof;
   117 
   118 val default_type = fst o dest_test_params o test_params_of
   119 
   120 val expect = snd o dest_test_params o test_params_of
   121 
   122 val map_test_params = Data.map o apsnd o map_test_params'
   123 
   124 val add_generator = Data.map o apfst o AList.update (op =);
   125 
   126 (* generating tests *)
   127 
   128 fun mk_tester ctxt t =
   129   let
   130     val name = Config.get ctxt tester
   131     val tester = case AList.lookup (op =) ((fst o Data.get o Context.Proof) ctxt) name
   132       of NONE => error ("No such quickcheck tester: " ^ name)
   133       | SOME tester => tester ctxt;
   134   in
   135     if Config.get ctxt quiet then
   136       try tester t
   137     else
   138       let
   139         val tester = Exn.interruptible_capture tester t
   140       in case Exn.get_result tester of
   141           NONE => SOME (Exn.release tester)
   142         | SOME tester => SOME tester
   143       end
   144   end
   145 
   146 (* testing propositions *)
   147 
   148 fun prep_test_term t =
   149   let
   150     val _ = (null (Term.add_tvars t []) andalso null (Term.add_tfrees t [])) orelse
   151       error "Term to be tested contains type variables";
   152     val _ = null (Term.add_vars t []) orelse
   153       error "Term to be tested contains schematic variables";
   154     val frees = Term.add_frees t [];
   155   in (frees, list_abs_free (frees, t)) end
   156 
   157 fun cpu_time description f =
   158   let
   159     val start = start_timing ()
   160     val result = Exn.capture f ()
   161     val time = Time.toMilliseconds (#cpu (end_timing start))
   162   in (Exn.release result, (description, time)) end
   163 
   164 fun limit ctxt is_interactive f exc () =
   165   TimeLimit.timeLimit (seconds (Config.get ctxt timeout)) f ()
   166   handle TimeLimit.TimeOut =>
   167     if is_interactive then exc () else raise TimeLimit.TimeOut
   168 
   169 fun test_term ctxt is_interactive t =
   170   let
   171     val (names, t') = apfst (map fst) (prep_test_term t);
   172     val current_size = Unsynchronized.ref 0
   173     fun excipit s =
   174       "Quickcheck " ^ s ^ " while testing at size " ^ string_of_int (!current_size)
   175     val (test_fun, comp_time) = cpu_time "quickcheck compilation" (fn () => mk_tester ctxt t');
   176     fun with_size test_fun k reports =
   177       if k > Config.get ctxt size then
   178         (NONE, reports)
   179       else
   180         let
   181           val _ = if Config.get ctxt quiet then () else Output.urgent_message
   182             ("Test data size: " ^ string_of_int k)
   183           val _ = current_size := k
   184           val ((result, new_report), timing) =
   185             cpu_time ("size " ^ string_of_int k) (fn () => test_fun (k - 1))
   186           val reports = case new_report of NONE => reports | SOME rep => (k, rep) :: reports
   187         in
   188           case result of NONE => with_size test_fun (k + 1) reports | SOME q => (SOME q, reports)
   189         end;
   190   in
   191     case test_fun of NONE => (NONE, ([comp_time], NONE))
   192       | SOME test_fun =>
   193         limit ctxt is_interactive (fn () =>
   194           let
   195             val ((result, reports), exec_time) =
   196               cpu_time "quickcheck execution" (fn () => with_size test_fun 1 [])
   197           in
   198             (case result of NONE => NONE | SOME ts => SOME (names ~~ ts),
   199               ([exec_time, comp_time],
   200                if Config.get ctxt report andalso not (null reports) then SOME reports else NONE))
   201           end) (fn () => error (excipit "ran out of time")) ()
   202   end;
   203 
   204 (* FIXME: this function shows that many assumptions are made upon the generation *)
   205 (* In the end there is probably no generic quickcheck interface left... *)
   206 
   207 fun test_term_with_increasing_cardinality ctxt is_interactive ts =
   208   let
   209     val thy = ProofContext.theory_of ctxt
   210     val (freess, ts') = split_list (map prep_test_term ts)
   211     val Ts = map snd (hd freess)
   212     val (test_funs, comp_time) = cpu_time "quickcheck compilation"
   213       (fn () => map (mk_tester ctxt) ts')
   214     fun test_card_size (card, size) =
   215       (* FIXME: why decrement size by one? *)
   216       case fst (the (nth test_funs (card - 1)) (size - 1)) of
   217         SOME ts => SOME (map fst (nth freess (card - 1)) ~~ ts)
   218       | NONE => NONE
   219     val enumeration_card_size =
   220       if forall (fn T => Sign.of_sort thy (T,  ["Enum.enum"])) Ts then
   221         (* size does not matter *)
   222         map (rpair 0) (1 upto (length ts))
   223       else
   224         (* size does matter *)
   225         map_product pair (1 upto (length ts)) (1 upto (Config.get ctxt size))
   226         |> sort (fn ((c1, s1), (c2, s2)) => int_ord ((c1 + s1), (c2 + s2)))
   227     in
   228       if (forall is_none test_funs) then
   229         (NONE, ([comp_time], NONE))
   230       else if (forall is_some test_funs) then
   231         limit ctxt is_interactive (fn () =>
   232           (get_first test_card_size enumeration_card_size, ([comp_time], NONE)))
   233           (fn () => error "Quickcheck ran out of time") ()
   234       else
   235         error "Unexpectedly, testers of some cardinalities are executable but others are not"
   236     end
   237 
   238 fun get_finite_types ctxt =
   239   fst (chop (Config.get ctxt finite_type_size)
   240     (map (Type o rpair []) ["Enum.finite_1", "Enum.finite_2", "Enum.finite_3",
   241      "Enum.finite_4", "Enum.finite_5"]))
   242 
   243 exception WELLSORTED of string
   244 
   245 fun monomorphic_term thy insts default_T =
   246   let
   247     fun subst (T as TFree (v, S)) =
   248       let
   249         val T' = AList.lookup (op =) insts v
   250           |> the_default default_T
   251       in if Sign.of_sort thy (T', S) then T'
   252         else raise (WELLSORTED ("For instantiation with default_type " ^
   253           Syntax.string_of_typ_global thy default_T ^
   254           ":\n" ^ Syntax.string_of_typ_global thy T' ^
   255           " to be substituted for variable " ^
   256           Syntax.string_of_typ_global thy T ^ " does not have sort " ^
   257           Syntax.string_of_sort_global thy S))
   258       end
   259       | subst T = T;
   260   in (map_types o map_atyps) subst end;
   261 
   262 datatype wellsorted_error = Wellsorted_Error of string | Term of term
   263 
   264 fun test_goal_terms lthy is_interactive insts check_goals =
   265   let
   266     val thy = ProofContext.theory_of lthy
   267     val default_insts =
   268       if Config.get lthy finite_types then (get_finite_types lthy) else (default_type lthy)
   269     val inst_goals =
   270       map (fn check_goal =>
   271         if not (null (Term.add_tfree_names check_goal [])) then
   272           map (fn T =>
   273             (pair (SOME T) o Term o Object_Logic.atomize_term thy o monomorphic_term thy insts T)
   274               check_goal
   275               handle WELLSORTED s => (SOME T, Wellsorted_Error s)) default_insts
   276         else
   277           [(NONE, Term (Object_Logic.atomize_term thy check_goal))]) check_goals
   278     val error_msg =
   279       cat_lines
   280         (maps (map_filter (fn (_, Term t) => NONE | (_, Wellsorted_Error s) => SOME s)) inst_goals)
   281     fun is_wellsorted_term (T, Term t) = SOME (T, t)
   282       | is_wellsorted_term (_, Wellsorted_Error s) = NONE
   283     val correct_inst_goals =
   284       case map (map_filter is_wellsorted_term) inst_goals of
   285         [[]] => error error_msg
   286       | xs => xs
   287     val _ = if Config.get lthy quiet then () else warning error_msg
   288     fun collect_results f reports [] = (NONE, rev reports)
   289       | collect_results f reports (t :: ts) =
   290         case f t of
   291           (SOME res, report) => (SOME res, rev (report :: reports))
   292         | (NONE, report) => collect_results f (report :: reports) ts
   293     fun test_term' goal =
   294       case goal of
   295         [(NONE, t)] => test_term lthy is_interactive t
   296       | ts => test_term_with_increasing_cardinality lthy is_interactive (map snd ts)
   297   in
   298     if Config.get lthy finite_types then
   299       collect_results test_term' [] correct_inst_goals
   300     else
   301       collect_results (test_term lthy is_interactive) [] (maps (map snd) correct_inst_goals)
   302   end;
   303 
   304 fun test_goal is_interactive insts i state =
   305   let
   306     val lthy = Proof.context_of state;
   307     val thy = Proof.theory_of state;
   308     fun strip (Const ("all", _) $ Abs (_, _, t)) = strip t
   309       | strip t = t;
   310     val {goal = st, ...} = Proof.raw_goal state;
   311     val (gi, frees) = Logic.goal_params (prop_of st) i;
   312     val some_locale = case (Option.map #target o Named_Target.peek) lthy
   313      of NONE => NONE
   314       | SOME "" => NONE
   315       | SOME locale => SOME locale;
   316     val assms = if Config.get lthy no_assms then [] else case some_locale
   317      of NONE => Assumption.all_assms_of lthy
   318       | SOME locale => Assumption.local_assms_of lthy (Locale.init locale thy);
   319     val proto_goal = Logic.list_implies (map Thm.term_of assms, subst_bounds (frees, strip gi));
   320     val check_goals = case some_locale
   321      of NONE => [proto_goal]
   322       | SOME locale => map (fn (_, phi) => Morphism.term phi proto_goal)
   323         (Locale.registrations_of (Context.Theory thy) (*FIXME*) locale);
   324   in
   325     test_goal_terms lthy is_interactive insts check_goals
   326   end
   327 
   328 (* pretty printing *)
   329 
   330 fun tool_name auto = (if auto then "Auto " else "") ^ "Quickcheck"
   331 
   332 fun pretty_counterex ctxt auto NONE = Pretty.str (tool_name auto ^ " found no counterexample.")
   333   | pretty_counterex ctxt auto (SOME cex) =
   334       Pretty.chunks (Pretty.str (tool_name auto ^ " found a counterexample:\n") ::
   335         map (fn (s, t) =>
   336           Pretty.block [Pretty.str (s ^ " ="), Pretty.brk 1, Syntax.pretty_term ctxt t]) (rev cex));
   337 
   338 fun pretty_report (Report {iterations = iterations, raised_match_errors = raised_match_errors,
   339     satisfied_assms = satisfied_assms, positive_concl_tests = positive_concl_tests}) =
   340   let
   341     fun pretty_stat s i = Pretty.block ([Pretty.str (s ^ ": " ^ string_of_int i)])
   342   in
   343      ([pretty_stat "iterations" iterations,
   344      pretty_stat "match exceptions" raised_match_errors]
   345      @ map_index
   346        (fn (i, n) => pretty_stat ("satisfied " ^ string_of_int (i + 1) ^ ". assumption") n)
   347        satisfied_assms
   348      @ [pretty_stat "positive conclusion tests" positive_concl_tests])
   349   end
   350 
   351 fun pretty_reports ctxt (SOME reports) =
   352   Pretty.chunks (Pretty.fbrk :: Pretty.str "Quickcheck report:" ::
   353     maps (fn (size, report) =>
   354       Pretty.str ("size " ^ string_of_int size ^ ":") :: pretty_report report @ [Pretty.brk 1])
   355       (rev reports))
   356   | pretty_reports ctxt NONE = Pretty.str ""
   357 
   358 fun pretty_counterex_and_reports ctxt auto (cex, timing_and_reports) =
   359   Pretty.chunks (pretty_counterex ctxt auto cex ::
   360     map (pretty_reports ctxt) (map snd timing_and_reports))
   361 
   362 (* automatic testing *)
   363 
   364 fun auto_quickcheck state =
   365   let
   366     val ctxt = Proof.context_of state;
   367     val res =
   368       state
   369       |> Proof.map_context (Config.put report false #> Config.put quiet true)
   370       |> try (test_goal false [] 1);
   371   in
   372     case res of
   373       NONE => (false, state)
   374     | SOME (NONE, report) => (false, state)
   375     | SOME (cex, report) => (true, Proof.goal_message (K (Pretty.chunks [Pretty.str "",
   376         Pretty.mark Markup.hilite (pretty_counterex ctxt true cex)])) state)
   377   end
   378 
   379 val setup = Auto_Tools.register_tool (auto, auto_quickcheck)
   380   #> setup_config
   381 
   382 (* Isar commands *)
   383 
   384 fun read_nat s = case (Library.read_int o Symbol.explode) s
   385  of (k, []) => if k >= 0 then k
   386       else error ("Not a natural number: " ^ s)
   387   | (_, _ :: _) => error ("Not a natural number: " ^ s);
   388 
   389 fun read_bool "false" = false
   390   | read_bool "true" = true
   391   | read_bool s = error ("Not a Boolean value: " ^ s)
   392 
   393 fun read_real s =
   394   case (Real.fromString s) of
   395     SOME s => s
   396   | NONE => error ("Not a real number: " ^ s)
   397 
   398 fun read_expectation "no_expectation" = No_Expectation
   399   | read_expectation "no_counterexample" = No_Counterexample
   400   | read_expectation "counterexample" = Counterexample
   401   | read_expectation s = error ("Not an expectation value: " ^ s)
   402 
   403 fun valid_tester_name genctxt = AList.defined (op =) (fst (Data.get genctxt))
   404 
   405 fun parse_tester name genctxt =
   406   if valid_tester_name genctxt name then
   407     Config.put_generic tester name genctxt
   408   else
   409     error ("Unknown tester: " ^ name)
   410 
   411 fun parse_test_param ("tester", [arg]) = parse_tester arg
   412   | parse_test_param ("size", [arg]) = Config.put_generic size (read_nat arg)
   413   | parse_test_param ("iterations", [arg]) = Config.put_generic iterations (read_nat arg)
   414   | parse_test_param ("default_type", arg) = (fn gen_ctxt =>
   415     map_test_params
   416       ((apfst o K) (map (ProofContext.read_typ (Context.proof_of gen_ctxt)) arg)) gen_ctxt)
   417   | parse_test_param ("no_assms", [arg]) = Config.put_generic no_assms (read_bool arg)
   418   | parse_test_param ("expect", [arg]) = map_test_params ((apsnd o K) (read_expectation arg))
   419   | parse_test_param ("report", [arg]) = Config.put_generic report (read_bool arg)
   420   | parse_test_param ("quiet", [arg]) = Config.put_generic quiet (read_bool arg)
   421   | parse_test_param ("timeout", [arg]) = Config.put_generic timeout (read_real arg)
   422   | parse_test_param ("finite_types", [arg]) = Config.put_generic finite_types (read_bool arg)
   423   | parse_test_param ("finite_type_size", [arg]) =
   424     Config.put_generic finite_type_size (read_nat arg)
   425   | parse_test_param (name, _) = fn genctxt =>
   426     if valid_tester_name genctxt name then
   427       Config.put_generic tester name genctxt
   428     else error ("Unknown tester or test parameter: " ^ name);
   429 
   430 fun parse_test_param_inst (name, arg) (insts, ctxt) =
   431       case try (ProofContext.read_typ ctxt) name
   432        of SOME (TFree (v, _)) => (apfst o AList.update (op =))
   433               (v, ProofContext.read_typ ctxt (the_single arg)) (insts, ctxt)
   434         | _ => (apsnd o Context.proof_map o parse_test_param) (name, arg) (insts, ctxt);
   435 
   436 fun quickcheck_params_cmd args = Context.theory_map (fold parse_test_param args);
   437 
   438 fun gen_quickcheck args i state =
   439   state
   440   |> Proof.map_context_result (fn ctxt => fold parse_test_param_inst args ([], ctxt))
   441   |> (fn (insts, state') => test_goal true insts i state'
   442   |> tap (fn (SOME x, _) => if expect (Proof.context_of state') = No_Counterexample then
   443                error ("quickcheck expected to find no counterexample but found one") else ()
   444            | (NONE, _) => if expect (Proof.context_of state') = Counterexample then
   445                error ("quickcheck expected to find a counterexample but did not find one") else ()))
   446 
   447 fun quickcheck args i state = fst (gen_quickcheck args i state);
   448 
   449 fun quickcheck_cmd args i state =
   450   gen_quickcheck args i (Toplevel.proof_of state)
   451   |> Pretty.writeln o pretty_counterex_and_reports (Toplevel.context_of state) false;
   452 
   453 val parse_arg =
   454   Parse.name --
   455     (Scan.optional (Parse.$$$ "=" |--
   456       (((Parse.name || Parse.float_number) >> single) ||
   457         (Parse.$$$ "[" |-- Parse.list1 Parse.name --| Parse.$$$ "]"))) ["true"]);
   458 
   459 val parse_args =
   460   Parse.$$$ "[" |-- Parse.list1 parse_arg --| Parse.$$$ "]" || Scan.succeed [];
   461 
   462 val _ =
   463   Outer_Syntax.command "quickcheck_params" "set parameters for random testing" Keyword.thy_decl
   464     (parse_args >> (fn args => Toplevel.theory (quickcheck_params_cmd args)));
   465 
   466 val _ =
   467   Outer_Syntax.improper_command "quickcheck" "try to find counterexample for subgoal" Keyword.diag
   468     (parse_args -- Scan.optional Parse.nat 1
   469       >> (fn (args, i) => Toplevel.no_timing o Toplevel.keep (quickcheck_cmd args i)));
   470 
   471 end;
   472 
   473 
   474 val auto_quickcheck = Quickcheck.auto;