src/Tools/quickcheck.ML
changeset 37929 22e0797857e6
parent 37913 e85f5ad02a8f
child 37974 d9549f9da779
equal deleted inserted replaced
37921:1e846be00ddf 37929:22e0797857e6
    11   val auto: bool Unsynchronized.ref
    11   val auto: bool Unsynchronized.ref
    12   val timing : bool Unsynchronized.ref
    12   val timing : bool Unsynchronized.ref
    13   datatype report = Report of
    13   datatype report = Report of
    14     { iterations : int, raised_match_errors : int,
    14     { iterations : int, raised_match_errors : int,
    15       satisfied_assms : int list, positive_concl_tests : int }
    15       satisfied_assms : int list, positive_concl_tests : int }
       
    16   datatype expectation = No_Expectation | No_Counterexample | Counterexample;
    16   datatype test_params = Test_Params of
    17   datatype test_params = Test_Params of
    17   { size: int, iterations: int, default_type: typ list, no_assms: bool, report: bool, quiet : bool};
    18   { size: int, iterations: int, default_type: typ list, no_assms: bool,
       
    19     expect : expectation, report: bool, quiet : bool};
    18   val test_params_of: theory -> test_params 
    20   val test_params_of: theory -> test_params 
    19   val add_generator:
    21   val add_generator:
    20     string * (Proof.context -> bool -> term -> int -> term list option * (bool list * bool))
    22     string * (Proof.context -> bool -> term -> int -> term list option * (bool list * bool))
    21       -> theory -> theory
    23       -> theory -> theory
    22   (* testing terms and proof states *)
    24   (* testing terms and proof states *)
    63       satisfied_assms =
    65       satisfied_assms =
    64         map2 (fn b => fn s => if b then s + 1 else s) assms
    66         map2 (fn b => fn s => if b then s + 1 else s) assms
    65          (if null satisfied_assms then replicate (length assms) 0 else satisfied_assms),
    67          (if null satisfied_assms then replicate (length assms) 0 else satisfied_assms),
    66       positive_concl_tests = if concl then positive_concl_tests + 1 else positive_concl_tests}
    68       positive_concl_tests = if concl then positive_concl_tests + 1 else positive_concl_tests}
    67 
    69 
       
    70 (* expectation *)
       
    71 
       
    72 datatype expectation = No_Expectation | No_Counterexample | Counterexample; 
       
    73 
       
    74 fun merge_expectation (expect1, expect2) =
       
    75   if expect1 = expect2 then expect1 else No_Expectation
       
    76 
    68 (* quickcheck configuration -- default parameters, test generators *)
    77 (* quickcheck configuration -- default parameters, test generators *)
    69 
    78 
    70 datatype test_params = Test_Params of
    79 datatype test_params = Test_Params of
    71   { size: int, iterations: int, default_type: typ list, no_assms: bool, report: bool, quiet : bool};
    80   { size: int, iterations: int, default_type: typ list, no_assms: bool,
    72 
    81   expect : expectation, report: bool, quiet : bool};
    73 fun dest_test_params (Test_Params { size, iterations, default_type, no_assms, report, quiet }) =
    82 
    74   ((size, iterations), ((default_type, no_assms), (report, quiet)));
    83 fun dest_test_params (Test_Params { size, iterations, default_type, no_assms, expect, report, quiet }) =
    75 fun make_test_params ((size, iterations), ((default_type, no_assms), (report, quiet))) =
    84   ((size, iterations), ((default_type, no_assms), ((expect, report), quiet)));
       
    85 fun make_test_params ((size, iterations), ((default_type, no_assms), ((expect, report), quiet))) =
    76   Test_Params { size = size, iterations = iterations, default_type = default_type,
    86   Test_Params { size = size, iterations = iterations, default_type = default_type,
    77                 no_assms = no_assms, report = report, quiet = quiet };
    87                 no_assms = no_assms, expect = expect, report = report, quiet = quiet };
    78 fun map_test_params f (Test_Params { size, iterations, default_type, no_assms, report, quiet }) =
    88 fun map_test_params f (Test_Params { size, iterations, default_type, no_assms, expect, report, quiet }) =
    79   make_test_params (f ((size, iterations), ((default_type, no_assms), (report, quiet))));
    89   make_test_params (f ((size, iterations), ((default_type, no_assms), ((expect, report), quiet))));
    80 fun merge_test_params (Test_Params { size = size1, iterations = iterations1, default_type = default_type1,
    90 fun merge_test_params (Test_Params { size = size1, iterations = iterations1, default_type = default_type1,
    81                                      no_assms = no_assms1, report = report1, quiet = quiet1 },
    91                                      no_assms = no_assms1, expect = expect1, report = report1, quiet = quiet1 },
    82   Test_Params { size = size2, iterations = iterations2, default_type = default_type2,
    92   Test_Params { size = size2, iterations = iterations2, default_type = default_type2,
    83                 no_assms = no_assms2, report = report2, quiet = quiet2 }) =
    93                 no_assms = no_assms2, expect = expect2, report = report2, quiet = quiet2 }) =
    84   make_test_params ((Int.max (size1, size2), Int.max (iterations1, iterations2)),
    94   make_test_params ((Int.max (size1, size2), Int.max (iterations1, iterations2)),
    85     ((merge (op =) (default_type1, default_type2), no_assms1 orelse no_assms2),
    95     ((merge (op =) (default_type1, default_type2), no_assms1 orelse no_assms2),
    86     (report1 orelse report2, quiet1 orelse quiet2)));
    96     ((merge_expectation (expect1, expect2), report1 orelse report2), quiet1 orelse quiet2)));
    87 
    97 
    88 structure Data = Theory_Data
    98 structure Data = Theory_Data
    89 (
    99 (
    90   type T = (string * (Proof.context -> bool -> term -> int -> term list option * (bool list * bool))) list
   100   type T = (string * (Proof.context -> bool -> term -> int -> term list option * (bool list * bool))) list
    91     * test_params;
   101     * test_params;
    92   val empty = ([], Test_Params
   102   val empty = ([], Test_Params
    93     { size = 10, iterations = 100, default_type = [], no_assms = false, report = false, quiet = false});
   103     { size = 10, iterations = 100, default_type = [], no_assms = false,
       
   104       expect = No_Expectation, report = false, quiet = false});
    94   val extend = I;
   105   val extend = I;
    95   fun merge ((generators1, params1), (generators2, params2)) : T =
   106   fun merge ((generators1, params1), (generators2, params2)) : T =
    96     (AList.merge (op =) (K true) (generators1, generators2),
   107     (AList.merge (op =) (K true) (generators1, generators2),
    97       merge_test_params (params1, params2));
   108       merge_test_params (params1, params2));
    98 );
   109 );
   269     (false, state)
   280     (false, state)
   270   else
   281   else
   271     let
   282     let
   272       val ctxt = Proof.context_of state;
   283       val ctxt = Proof.context_of state;
   273       val assms = map term_of (Assumption.all_assms_of ctxt);
   284       val assms = map term_of (Assumption.all_assms_of ctxt);
   274       val Test_Params { size, iterations, default_type, no_assms, report, quiet } =
   285       val Test_Params {size, iterations, default_type, no_assms, ...} =
   275         (snd o Data.get o Proof.theory_of) state;
   286         (snd o Data.get o Proof.theory_of) state;
   276       val res =
   287       val res =
   277         try (test_goal true false NONE size iterations default_type no_assms [] 1 assms) state;
   288         try (test_goal true false NONE size iterations default_type no_assms [] 1 assms) state;
   278     in
   289     in
   279       case res of
   290       case res of
   295 
   306 
   296 fun read_bool "false" = false
   307 fun read_bool "false" = false
   297   | read_bool "true" = true
   308   | read_bool "true" = true
   298   | read_bool s = error ("Not a Boolean value: " ^ s)
   309   | read_bool s = error ("Not a Boolean value: " ^ s)
   299 
   310 
       
   311 fun read_expectation "no_expectation" = No_Expectation
       
   312   | read_expectation "no_counterexample" = No_Counterexample 
       
   313   | read_expectation "counterexample" = Counterexample
       
   314   | read_expectation s = error ("Not an expectation value: " ^ s)  
       
   315 
   300 fun parse_test_param ctxt ("size", [arg]) =
   316 fun parse_test_param ctxt ("size", [arg]) =
   301       (apfst o apfst o K) (read_nat arg)
   317       (apfst o apfst o K) (read_nat arg)
   302   | parse_test_param ctxt ("iterations", [arg]) =
   318   | parse_test_param ctxt ("iterations", [arg]) =
   303       (apfst o apsnd o K) (read_nat arg)
   319       (apfst o apsnd o K) (read_nat arg)
   304   | parse_test_param ctxt ("default_type", arg) =
   320   | parse_test_param ctxt ("default_type", arg) =
   305       (apsnd o apfst o apfst o K) (map (ProofContext.read_typ ctxt) arg)
   321       (apsnd o apfst o apfst o K) (map (ProofContext.read_typ ctxt) arg)
   306   | parse_test_param ctxt ("no_assms", [arg]) =
   322   | parse_test_param ctxt ("no_assms", [arg]) =
   307       (apsnd o apfst o apsnd o K) (read_bool arg)
   323       (apsnd o apfst o apsnd o K) (read_bool arg)
       
   324   | parse_test_param ctxt ("expect", [arg]) =
       
   325       (apsnd o apsnd o apfst o apfst o K) (read_expectation arg)
   308   | parse_test_param ctxt ("report", [arg]) =
   326   | parse_test_param ctxt ("report", [arg]) =
   309       (apsnd o apsnd o apfst o K) (read_bool arg)
   327       (apsnd o apsnd o apfst o apsnd o K) (read_bool arg)
   310   | parse_test_param ctxt ("quiet", [arg]) =
   328   | parse_test_param ctxt ("quiet", [arg]) =
   311       (apsnd o apsnd o apsnd o K) (read_bool arg)
   329       (apsnd o apsnd o apsnd o K) (read_bool arg)       
   312   | parse_test_param ctxt (name, _) =
   330   | parse_test_param ctxt (name, _) =
   313       error ("Unknown test parameter: " ^ name);
   331       error ("Unknown test parameter: " ^ name);
   314 
   332 
   315 fun parse_test_param_inst ctxt ("generator", [arg]) =
   333 fun parse_test_param_inst ctxt ("generator", [arg]) =
   316       (apsnd o apfst o K o SOME) arg
   334       (apsnd o apfst o K o SOME) arg
   334     val thy = Proof.theory_of state;
   352     val thy = Proof.theory_of state;
   335     val ctxt = Proof.context_of state;
   353     val ctxt = Proof.context_of state;
   336     val assms = map term_of (Assumption.all_assms_of ctxt);
   354     val assms = map term_of (Assumption.all_assms_of ctxt);
   337     val default_params = (dest_test_params o snd o Data.get) thy;
   355     val default_params = (dest_test_params o snd o Data.get) thy;
   338     val f = fold (parse_test_param_inst ctxt) args;
   356     val f = fold (parse_test_param_inst ctxt) args;
   339     val (((size, iterations), ((default_type, no_assms), (report, quiet))), (generator_name, insts)) =
   357     val (((size, iterations), ((default_type, no_assms), ((expect, report), quiet))), (generator_name, insts)) =
   340       f (default_params, (NONE, []));
   358       f (default_params, (NONE, []));
   341   in
   359   in
   342     test_goal quiet report generator_name size iterations default_type no_assms insts i assms state
   360     test_goal quiet report generator_name size iterations default_type no_assms insts i assms state
       
   361     |> tap (fn (SOME x, _) => if expect = No_Counterexample then
       
   362                  error ("quickcheck expect to find no counterexample but found one") else ()
       
   363              | (NONE, _) => if expect = Counterexample then
       
   364                  error ("quickcheck expected to find a counterexample but did not find one") else ())
   343   end;
   365   end;
   344 
   366 
   345 fun quickcheck args i state = fst (gen_quickcheck args i state);
   367 fun quickcheck args i state = fst (gen_quickcheck args i state);
   346 
   368 
   347 fun quickcheck_cmd args i state =
   369 fun quickcheck_cmd args i state =