wenzelm@30826: (* Title: Tools/quickcheck.ML bulwahn@40481: Author: Stefan Berghofer, Florian Haftmann, Lukas Bulwahn, TU Muenchen haftmann@28256: haftmann@28256: Generic counterexample search engine. haftmann@28256: *) haftmann@28256: haftmann@28256: signature QUICKCHECK = haftmann@28256: sig blanchet@43861: val quickcheckN: string blanchet@43861: val genuineN: string blanchet@43861: val noneN: string blanchet@43861: val unknownN: string bulwahn@38150: val setup: theory -> theory bulwahn@38150: (* configuration *) wenzelm@32740: val auto: bool Unsynchronized.ref bulwahn@44753: val batch_tester : string Config.T bulwahn@40892: val size : int Config.T bulwahn@40892: val iterations : int Config.T bulwahn@46084: val depth : int Config.T bulwahn@40892: val no_assms : bool Config.T bulwahn@40892: val report : bool Config.T bulwahn@42952: val timing : bool Config.T bulwahn@40892: val quiet : bool Config.T bulwahn@40892: val timeout : real Config.T bulwahn@46320: val allow_function_inversion : bool Config.T; bulwahn@40896: val finite_types : bool Config.T bulwahn@40896: val finite_type_size : int Config.T bulwahn@44783: val set_active_testers: string list -> Context.generic -> Context.generic wenzelm@41765: datatype expectation = No_Expectation | No_Counterexample | Counterexample; bulwahn@40892: datatype test_params = Test_Params of {default_type: typ list, expect : expectation}; bulwahn@40481: val test_params_of : Proof.context -> test_params bulwahn@40892: val map_test_params : (typ list * expectation -> typ list * expectation) bulwahn@40481: -> Context.generic -> Context.generic bulwahn@46029: val default_type : Proof.context -> typ list bulwahn@42953: datatype report = Report of bulwahn@42953: { iterations : int, raised_match_errors : int, bulwahn@42953: satisfied_assms : int list, positive_concl_tests : int } bulwahn@42953: (* quickcheck's result *) bulwahn@42952: datatype result = bulwahn@42952: Result of bulwahn@42952: {counterexample : (string * term) list option, bulwahn@42952: evaluation_terms : (term * term) list option, bulwahn@42952: timings : (string * int) list, bulwahn@42952: reports : (int * report) list} bulwahn@44185: val empty_result : result bulwahn@46029: val found_counterexample : result -> bool bulwahn@44456: val add_timing : (string * int) -> result Unsynchronized.ref -> unit bulwahn@46029: val add_response : string list -> term list -> term list option -> result Unsynchronized.ref -> unit bulwahn@46029: val add_report : int -> report option -> result Unsynchronized.ref -> unit bulwahn@42953: val counterexample_of : result -> (string * term) list option bulwahn@42953: val timings_of : result -> (string * int) list bulwahn@44749: (* registering testers & generators *) bulwahn@44749: type tester = bulwahn@46290: Proof.context -> bool -> (string * typ) list -> (term * term list) list -> result list bulwahn@44749: val add_tester : string * (bool Config.T * tester) -> Context.generic -> Context.generic bulwahn@44749: val add_batch_generator : bulwahn@43953: string * (Proof.context -> term list -> (int -> term list option) list) bulwahn@43953: -> Context.generic -> Context.generic bulwahn@44749: val add_batch_validator : bulwahn@43953: string * (Proof.context -> term list -> (int -> bool) list) bulwahn@43953: -> Context.generic -> Context.generic bulwahn@43954: (* basic operations *) bulwahn@46029: val message : Proof.context -> string -> unit bulwahn@46029: val limit : Time.time -> (bool * bool) -> (unit -> 'a) -> (unit -> 'a) -> unit -> 'a bulwahn@38150: (* testing terms and proof states *) bulwahn@46029: val mk_batch_validator : Proof.context -> term list -> (int -> bool) list option bulwahn@46029: val mk_batch_tester : Proof.context -> term list -> (int -> term list option) list option bulwahn@46029: val active_testers : Proof.context -> tester list bulwahn@46029: val test_terms : bulwahn@44754: Proof.context -> bool * bool -> (string * typ) list -> (term * term list) list -> result list option bulwahn@38149: val quickcheck: (string * string list) list -> int -> Proof.state -> (string * term) list option haftmann@28256: end; haftmann@28256: haftmann@28256: structure Quickcheck : QUICKCHECK = haftmann@28256: struct haftmann@28256: blanchet@43861: val quickcheckN = "quickcheck" blanchet@43861: val quickcheck_paramsN = "quickcheck_params" blanchet@43861: blanchet@43861: val genuineN = "genuine" blanchet@43861: val noneN = "none" blanchet@43861: val unknownN = "unknown" blanchet@43861: wenzelm@30981: (* preferences *) wenzelm@30981: wenzelm@32740: val auto = Unsynchronized.ref false; wenzelm@30981: wenzelm@30981: val _ = wenzelm@30981: ProofGeneralPgip.add_preference Preferences.category_tracing wenzelm@39862: (Unsynchronized.setmp auto true (fn () => wenzelm@30981: Preferences.bool_pref auto wenzelm@30981: "auto-quickcheck" blanchet@39575: "Run Quickcheck automatically.") ()); wenzelm@30981: bulwahn@35378: (* quickcheck report *) bulwahn@35378: bulwahn@35378: datatype report = Report of bulwahn@35378: { iterations : int, raised_match_errors : int, bulwahn@35378: satisfied_assms : int list, positive_concl_tests : int } bulwahn@35378: bulwahn@42952: (* Quickcheck Result *) bulwahn@42952: bulwahn@42952: datatype result = Result of bulwahn@42952: { counterexample : (string * term) list option, evaluation_terms : (term * term) list option, bulwahn@42952: timings : (string * int) list, reports : (int * report) list} bulwahn@42952: bulwahn@42952: val empty_result = bulwahn@42952: Result {counterexample = NONE, evaluation_terms = NONE, timings = [], reports = []} bulwahn@42952: bulwahn@42952: fun counterexample_of (Result r) = #counterexample r bulwahn@42952: bulwahn@42952: fun found_counterexample (Result r) = is_some (#counterexample r) bulwahn@42952: bulwahn@42952: fun response_of (Result r) = case (#counterexample r, #evaluation_terms r) of bulwahn@42952: (SOME ts, SOME evals) => SOME (ts, evals) bulwahn@42952: | (NONE, NONE) => NONE bulwahn@42952: bulwahn@42952: fun timings_of (Result r) = #timings r bulwahn@42952: bulwahn@46029: fun set_response names eval_terms (SOME ts) (Result r) = bulwahn@42952: let bulwahn@42952: val (ts1, ts2) = chop (length names) ts bulwahn@46029: val (eval_terms', _) = chop (length ts2) eval_terms bulwahn@42952: in bulwahn@46029: Result {counterexample = SOME (names ~~ ts1), evaluation_terms = SOME (eval_terms' ~~ ts2), bulwahn@42952: timings = #timings r, reports = #reports r} bulwahn@42952: end bulwahn@46029: | set_response _ _ NONE result = result bulwahn@42952: bulwahn@44456: bulwahn@44456: fun set_counterexample counterexample (Result r) = bulwahn@44456: Result {counterexample = counterexample, evaluation_terms = #evaluation_terms r, bulwahn@44456: timings = #timings r, reports = #reports r} bulwahn@44456: bulwahn@44456: fun set_evaluation_terms evaluation_terms (Result r) = bulwahn@44456: Result {counterexample = #counterexample r, evaluation_terms = evaluation_terms, bulwahn@44456: timings = #timings r, reports = #reports r} bulwahn@44456: bulwahn@42952: fun cons_timing timing (Result r) = bulwahn@42952: Result {counterexample = #counterexample r, evaluation_terms = #evaluation_terms r, bulwahn@42952: timings = cons timing (#timings r), reports = #reports r} bulwahn@42952: bulwahn@42952: fun cons_report size (SOME report) (Result r) = bulwahn@42952: Result {counterexample = #counterexample r, evaluation_terms = #evaluation_terms r, bulwahn@42952: timings = #timings r, reports = cons (size, report) (#reports r)} bulwahn@42952: | cons_report _ NONE result = result bulwahn@42952: wenzelm@43071: fun add_timing timing result_ref = wenzelm@43071: Unsynchronized.change result_ref (cons_timing timing) bulwahn@42952: wenzelm@43071: fun add_report size report result_ref = wenzelm@43071: Unsynchronized.change result_ref (cons_report size report) bulwahn@42952: bulwahn@42952: fun add_response names eval_terms response result_ref = bulwahn@46029: Unsynchronized.change result_ref (set_response names eval_terms response) bulwahn@42952: bulwahn@38169: (* expectation *) bulwahn@38169: wenzelm@41765: datatype expectation = No_Expectation | No_Counterexample | Counterexample; bulwahn@38169: bulwahn@38169: fun merge_expectation (expect1, expect2) = bulwahn@38169: if expect1 = expect2 then expect1 else No_Expectation bulwahn@38169: haftmann@28315: (* quickcheck configuration -- default parameters, test generators *) bulwahn@44753: val batch_tester = Attrib.setup_config_string @{binding quickcheck_batch_tester} (K "") wenzelm@43487: val size = Attrib.setup_config_int @{binding quickcheck_size} (K 10) wenzelm@43487: val iterations = Attrib.setup_config_int @{binding quickcheck_iterations} (K 100) bulwahn@46084: val depth = Attrib.setup_config_int @{binding quickcheck_depth} (K 10) bulwahn@46084: wenzelm@43487: val no_assms = Attrib.setup_config_bool @{binding quickcheck_no_assms} (K false) wenzelm@43487: val report = Attrib.setup_config_bool @{binding quickcheck_report} (K true) wenzelm@43487: val timing = Attrib.setup_config_bool @{binding quickcheck_timing} (K false) wenzelm@43487: val quiet = Attrib.setup_config_bool @{binding quickcheck_quiet} (K false) wenzelm@43487: val timeout = Attrib.setup_config_real @{binding quickcheck_timeout} (K 30.0) bulwahn@46320: val allow_function_inversion = bulwahn@46320: Attrib.setup_config_bool @{binding quickcheck_allow_function_inversion} (K false) wenzelm@43487: val finite_types = Attrib.setup_config_bool @{binding quickcheck_finite_types} (K true) wenzelm@43487: val finite_type_size = Attrib.setup_config_int @{binding quickcheck_finite_type_size} (K 3) bulwahn@40894: haftmann@28309: datatype test_params = Test_Params of bulwahn@40892: {default_type: typ list, expect : expectation}; haftmann@28309: bulwahn@40892: fun dest_test_params (Test_Params {default_type, expect}) = (default_type, expect); wenzelm@39034: wenzelm@41765: fun make_test_params (default_type, expect) = wenzelm@41765: Test_Params {default_type = default_type, expect = expect}; wenzelm@39034: wenzelm@41765: fun map_test_params' f (Test_Params {default_type, expect}) = wenzelm@41765: make_test_params (f (default_type, expect)); wenzelm@39034: wenzelm@39034: fun merge_test_params wenzelm@41720: (Test_Params {default_type = default_type1, expect = expect1}, wenzelm@41720: Test_Params {default_type = default_type2, expect = expect2}) = wenzelm@41720: make_test_params wenzelm@41720: (merge (op =) (default_type1, default_type2), merge_expectation (expect1, expect2)); haftmann@28309: bulwahn@44749: type tester = bulwahn@46290: Proof.context -> bool -> (string * typ) list -> (term * term list) list -> result list bulwahn@44749: bulwahn@39479: structure Data = Generic_Data wenzelm@33522: ( wenzelm@39034: type T = bulwahn@44749: ((string * (bool Config.T * tester)) list bulwahn@43065: * ((string * (Proof.context -> term list -> (int -> term list option) list)) list bulwahn@43065: * ((string * (Proof.context -> term list -> (int -> bool) list)) list))) wenzelm@39034: * test_params; bulwahn@44748: val empty = (([], ([], [])), Test_Params {default_type = [], expect = No_Expectation}); haftmann@28256: val extend = I; bulwahn@44748: fun merge (((testers1, (batch_generators1, batch_validators1)), params1), bulwahn@44748: ((testers2, (batch_generators2, batch_validators2)), params2)) : T = bulwahn@44748: ((AList.merge (op =) (K true) (testers1, testers2), bulwahn@43953: (AList.merge (op =) (K true) (batch_generators1, batch_generators2), bulwahn@43953: AList.merge (op =) (K true) (batch_validators1, batch_validators2))), haftmann@28309: merge_test_params (params1, params2)); wenzelm@33522: ); haftmann@28256: bulwahn@39479: val test_params_of = snd o Data.get o Context.Proof; bulwahn@38150: bulwahn@40892: val default_type = fst o dest_test_params o test_params_of bulwahn@40481: bulwahn@40892: val expect = snd o dest_test_params o test_params_of bulwahn@40481: bulwahn@40481: val map_test_params = Data.map o apsnd o map_test_params' bulwahn@39480: bulwahn@44748: val add_tester = Data.map o apfst o apfst o AList.update (op =); bulwahn@42733: bulwahn@43065: val add_batch_generator = Data.map o apfst o apsnd o apfst o AList.update (op =); bulwahn@43065: bulwahn@43065: val add_batch_validator = Data.map o apfst o apsnd o apsnd o AList.update (op =); haftmann@28256: bulwahn@44752: fun active_testers ctxt = bulwahn@44752: let bulwahn@44752: val testers = (map snd o fst o fst o Data.get o Context.Proof) ctxt bulwahn@44752: in bulwahn@44752: map snd (filter (fn (active, _) => Config.get ctxt active) testers) bulwahn@44752: end bulwahn@44752: bulwahn@44755: fun set_active_testers [] gen_ctxt = gen_ctxt bulwahn@44755: | set_active_testers testers gen_ctxt = bulwahn@44752: let bulwahn@44752: val registered_testers = (fst o fst o Data.get) gen_ctxt bulwahn@44752: in bulwahn@44752: fold (fn (name, (config, _)) => Config.put_generic config (member (op =) testers name)) bulwahn@44752: registered_testers gen_ctxt bulwahn@44752: end bulwahn@44752: haftmann@28315: (* generating tests *) haftmann@28315: bulwahn@42733: fun gen_mk_tester lookup ctxt v = haftmann@28309: let bulwahn@44753: val name = Config.get ctxt batch_tester bulwahn@42733: val tester = case lookup ctxt name bulwahn@44753: of NONE => error ("No such quickcheck batch-tester: " ^ name) bulwahn@41153: | SOME tester => tester ctxt; wenzelm@40491: in bulwahn@41153: if Config.get ctxt quiet then bulwahn@42733: try tester v bulwahn@41153: else wenzelm@44633: let (* FIXME !?!? *) bulwahn@42733: val tester = Exn.interruptible_capture tester v wenzelm@44633: in case Exn.get_res tester of bulwahn@41153: NONE => SOME (Exn.release tester) bulwahn@41153: | SOME tester => SOME tester bulwahn@41153: end bulwahn@41153: end bulwahn@44753: bulwahn@43065: val mk_batch_tester = bulwahn@43065: gen_mk_tester (fn ctxt => AList.lookup (op =) ((fst o snd o fst o Data.get o Context.Proof) ctxt)) bulwahn@43065: val mk_batch_validator = bulwahn@43065: gen_mk_tester (fn ctxt => AList.lookup (op =) ((snd o snd o fst o Data.get o Context.Proof) ctxt)) bulwahn@43953: bulwahn@44748: fun lookup_tester ctxt = AList.lookup (op =) ((fst o fst o Data.get o Context.Proof) ctxt) bulwahn@43953: haftmann@28315: (* testing propositions *) haftmann@28315: bulwahn@44747: type compile_generator = bulwahn@44747: Proof.context -> (term * term list) list -> int list -> term list option * report option bulwahn@44747: bulwahn@44456: fun limit timeout (limit_time, is_interactive) f exc () = bulwahn@42625: if limit_time then bulwahn@44456: TimeLimit.timeLimit timeout f () bulwahn@42625: handle TimeLimit.TimeOut => bulwahn@42625: if is_interactive then exc () else raise TimeLimit.TimeOut bulwahn@42625: else bulwahn@42625: f () bulwahn@42624: bulwahn@44786: fun message ctxt s = if Config.get ctxt quiet then () else Output.urgent_message s bulwahn@44749: bulwahn@46029: fun test_terms ctxt (limit_time, is_interactive) insts goals = bulwahn@44749: case active_testers ctxt of bulwahn@44752: [] => error "No active testers for quickcheck" bulwahn@44786: | testers => limit (seconds (Config.get ctxt timeout)) (limit_time, is_interactive) bulwahn@44786: (fn () => Par_List.get_some (fn tester => bulwahn@46290: tester ctxt (length testers > 1) insts goals |> bulwahn@44786: (fn result => if exists found_counterexample result then SOME result else NONE)) testers) bulwahn@44786: (fn () => (message ctxt "Quickcheck ran out of time"; NONE)) (); bulwahn@44786: bulwahn@42897: fun test_goal (time_limit, is_interactive) (insts, eval_terms) i state = bulwahn@40896: let bulwahn@40896: val lthy = Proof.context_of state; bulwahn@40896: val thy = Proof.theory_of state; bulwahn@40896: fun strip (Const ("all", _) $ Abs (_, _, t)) = strip t bulwahn@40896: | strip t = t; bulwahn@40896: val {goal = st, ...} = Proof.raw_goal state; bulwahn@40896: val (gi, frees) = Logic.goal_params (prop_of st) i; bulwahn@40896: val some_locale = case (Option.map #target o Named_Target.peek) lthy bulwahn@40896: of NONE => NONE bulwahn@40896: | SOME "" => NONE bulwahn@40896: | SOME locale => SOME locale; bulwahn@40896: val assms = if Config.get lthy no_assms then [] else case some_locale bulwahn@40896: of NONE => Assumption.all_assms_of lthy bulwahn@40896: | SOME locale => Assumption.local_assms_of lthy (Locale.init locale thy); bulwahn@40896: val proto_goal = Logic.list_implies (map Thm.term_of assms, subst_bounds (frees, strip gi)); bulwahn@43954: val goals = case some_locale bulwahn@42897: of NONE => [(proto_goal, eval_terms)] bulwahn@42897: | SOME locale => bulwahn@42897: map (fn (_, phi) => (Morphism.term phi proto_goal, map (Morphism.term phi) eval_terms)) bulwahn@42897: (Locale.registrations_of (Context.Theory thy) (*FIXME*) locale); bulwahn@40896: in bulwahn@46029: test_terms lthy (time_limit, is_interactive) insts goals bulwahn@40896: end bulwahn@40896: bulwahn@38152: (* pretty printing *) haftmann@28315: blanchet@40466: fun tool_name auto = (if auto then "Auto " else "") ^ "Quickcheck" blanchet@40466: blanchet@40466: fun pretty_counterex ctxt auto NONE = Pretty.str (tool_name auto ^ " found no counterexample.") bulwahn@42899: | pretty_counterex ctxt auto (SOME (cex, eval_terms)) = bulwahn@42899: Pretty.chunks ((Pretty.str (tool_name auto ^ " found a counterexample:\n") :: haftmann@28315: map (fn (s, t) => bulwahn@42899: Pretty.block [Pretty.str (s ^ " ="), Pretty.brk 1, Syntax.pretty_term ctxt t]) (rev cex)) blanchet@43861: @ (if null eval_terms then [] bulwahn@46029: else (Pretty.fbrk :: Pretty.str ("Evaluated terms:") :: blanchet@43861: map (fn (t, u) => bulwahn@44752: Pretty.block [Syntax.pretty_term ctxt t, Pretty.str " =", Pretty.brk 1, bulwahn@44752: Syntax.pretty_term ctxt u]) (rev eval_terms)))); haftmann@28315: bulwahn@35378: fun pretty_report (Report {iterations = iterations, raised_match_errors = raised_match_errors, bulwahn@35378: satisfied_assms = satisfied_assms, positive_concl_tests = positive_concl_tests}) = bulwahn@35378: let bulwahn@35378: fun pretty_stat s i = Pretty.block ([Pretty.str (s ^ ": " ^ string_of_int i)]) bulwahn@35378: in bulwahn@35378: ([pretty_stat "iterations" iterations, bulwahn@35378: pretty_stat "match exceptions" raised_match_errors] wenzelm@41765: @ map_index wenzelm@41765: (fn (i, n) => pretty_stat ("satisfied " ^ string_of_int (i + 1) ^ ". assumption") n) bulwahn@35378: satisfied_assms bulwahn@35378: @ [pretty_stat "positive conclusion tests" positive_concl_tests]) bulwahn@35378: end bulwahn@35378: bulwahn@35380: fun pretty_reports ctxt (SOME reports) = bulwahn@41160: Pretty.chunks (Pretty.fbrk :: Pretty.str "Quickcheck report:" :: bulwahn@41153: maps (fn (size, report) => bulwahn@41153: Pretty.str ("size " ^ string_of_int size ^ ":") :: pretty_report report @ [Pretty.brk 1]) bulwahn@35378: (rev reports)) bulwahn@35380: | pretty_reports ctxt NONE = Pretty.str "" bulwahn@35378: bulwahn@42952: fun pretty_timings timings = bulwahn@42952: Pretty.chunks (Pretty.fbrk :: Pretty.str "timings:" :: bulwahn@42952: maps (fn (label, time) => bulwahn@42952: Pretty.str (label ^ ": " ^ string_of_int time ^ " ms") :: Pretty.brk 1 :: []) (rev timings)) bulwahn@42952: bulwahn@43299: fun pretty_counterex_and_reports ctxt auto [] = bulwahn@43299: Pretty.chunks [Pretty.strs (tool_name auto :: bulwahn@43299: space_explode " " "is used in a locale where no interpretation is provided."), bulwahn@43299: Pretty.strs (space_explode " " "Please provide an executable interpretation for the locale.")] bulwahn@43299: | pretty_counterex_and_reports ctxt auto (result :: _) = bulwahn@43299: Pretty.chunks (pretty_counterex ctxt auto (response_of result) :: bulwahn@43299: (* map (pretty_reports ctxt) (reports_of result) :: *) bulwahn@43299: (if Config.get ctxt timing then cons (pretty_timings (timings_of result)) else I) []) haftmann@28315: wenzelm@30981: (* Isar commands *) haftmann@28315: haftmann@28336: fun read_nat s = case (Library.read_int o Symbol.explode) s haftmann@28336: of (k, []) => if k >= 0 then k haftmann@28336: else error ("Not a natural number: " ^ s) haftmann@28336: | (_, _ :: _) => error ("Not a natural number: " ^ s); bulwahn@38149: blanchet@34125: fun read_bool "false" = false blanchet@34125: | read_bool "true" = true blanchet@34125: | read_bool s = error ("Not a Boolean value: " ^ s) haftmann@28315: bulwahn@40612: fun read_real s = bulwahn@40612: case (Real.fromString s) of bulwahn@40612: SOME s => s bulwahn@40612: | NONE => error ("Not a real number: " ^ s) bulwahn@40612: bulwahn@38169: fun read_expectation "no_expectation" = No_Expectation wenzelm@41765: | read_expectation "no_counterexample" = No_Counterexample bulwahn@38169: | read_expectation "counterexample" = Counterexample wenzelm@41765: | read_expectation s = error ("Not an expectation value: " ^ s) bulwahn@38169: bulwahn@44748: fun valid_tester_name genctxt name = AList.defined (op =) (fst (fst (Data.get genctxt))) name bulwahn@44748: bulwahn@44752: fun parse_tester name (testers, genctxt) = bulwahn@41156: if valid_tester_name genctxt name then bulwahn@44752: (insert (op =) name testers, genctxt) bulwahn@41156: else bulwahn@41156: error ("Unknown tester: " ^ name) bulwahn@41156: bulwahn@44752: fun parse_test_param ("tester", args) = fold parse_tester args bulwahn@44752: | parse_test_param ("size", [arg]) = apsnd (Config.put_generic size (read_nat arg)) bulwahn@44752: | parse_test_param ("iterations", [arg]) = apsnd (Config.put_generic iterations (read_nat arg)) bulwahn@46084: | parse_test_param ("depth", [arg]) = apsnd (Config.put_generic depth (read_nat arg)) bulwahn@44752: | parse_test_param ("default_type", arg) = (fn (testers, gen_ctxt) => bulwahn@44752: (testers, map_test_params bulwahn@44752: ((apfst o K) (map (Proof_Context.read_typ (Context.proof_of gen_ctxt)) arg)) gen_ctxt)) bulwahn@44752: | parse_test_param ("no_assms", [arg]) = apsnd (Config.put_generic no_assms (read_bool arg)) bulwahn@44752: | parse_test_param ("expect", [arg]) = apsnd (map_test_params ((apsnd o K) (read_expectation arg))) bulwahn@44752: | parse_test_param ("report", [arg]) = apsnd (Config.put_generic report (read_bool arg)) bulwahn@44752: | parse_test_param ("quiet", [arg]) = apsnd (Config.put_generic quiet (read_bool arg)) bulwahn@44752: | parse_test_param ("timeout", [arg]) = apsnd (Config.put_generic timeout (read_real arg)) bulwahn@44752: | parse_test_param ("finite_types", [arg]) = apsnd (Config.put_generic finite_types (read_bool arg)) bulwahn@46320: | parse_test_param ("allow_function_inversion", [arg]) = bulwahn@46320: apsnd (Config.put_generic allow_function_inversion (read_bool arg)) wenzelm@41765: | parse_test_param ("finite_type_size", [arg]) = bulwahn@44752: apsnd (Config.put_generic finite_type_size (read_nat arg)) bulwahn@44752: | parse_test_param (name, _) = (fn (testers, genctxt) => bulwahn@41156: if valid_tester_name genctxt name then bulwahn@44752: (insert (op =) name testers, genctxt) bulwahn@44752: else error ("Unknown tester or test parameter: " ^ name)); haftmann@28315: bulwahn@44752: fun proof_map_result f = apsnd Context.the_proof o f o Context.Proof; bulwahn@44752: bulwahn@44752: fun parse_test_param_inst (name, arg) ((insts, eval_terms), (testers, ctxt)) = wenzelm@43232: case try (Proof_Context.read_typ ctxt) name bulwahn@42896: of SOME (TFree (v, _)) => bulwahn@44752: ((AList.update (op =) (v, Proof_Context.read_typ ctxt (the_single arg)) insts, eval_terms), bulwahn@44752: (testers, ctxt)) bulwahn@42896: | NONE => (case name of bulwahn@44752: "eval" => ((insts, eval_terms @ map (Syntax.read_term ctxt) arg), (testers, ctxt)) bulwahn@44752: | _ => ((insts, eval_terms), bulwahn@44752: proof_map_result (fn gen_ctxt => parse_test_param (name, arg) (testers, gen_ctxt)) ctxt)); haftmann@28315: bulwahn@44752: fun quickcheck_params_cmd args = Context.theory_map bulwahn@44752: (fn gen_ctxt => uncurry set_active_testers (fold parse_test_param args ([], gen_ctxt))); wenzelm@41765: bulwahn@42952: fun check_expectation state results = bulwahn@44750: (if is_some results andalso expect (Proof.context_of state) = No_Counterexample bulwahn@42952: then bulwahn@42952: error ("quickcheck expected to find no counterexample but found one") bulwahn@42952: else bulwahn@44750: (if is_none results andalso expect (Proof.context_of state) = Counterexample bulwahn@42952: then bulwahn@42952: error ("quickcheck expected to find a counterexample but did not find one") bulwahn@42952: else ())) bulwahn@42952: bulwahn@35378: fun gen_quickcheck args i state = bulwahn@40892: state bulwahn@44752: |> Proof.map_context_result (fn ctxt => bulwahn@44752: apsnd (fn (testers, ctxt) => Context.proof_map (set_active_testers testers) ctxt) bulwahn@44752: (fold parse_test_param_inst args (([], []), ([], ctxt)))) bulwahn@42897: |> (fn ((insts, eval_terms), state') => test_goal (true, true) (insts, eval_terms) i state' bulwahn@42952: |> tap (check_expectation state')) boehmes@32295: bulwahn@44750: fun quickcheck args i state = bulwahn@44750: Option.map (the o get_first counterexample_of) (gen_quickcheck args i state) bulwahn@35378: boehmes@32295: fun quickcheck_cmd args i state = bulwahn@35378: gen_quickcheck args i (Toplevel.proof_of state) bulwahn@44750: |> Option.map (the o get_first response_of) blanchet@43861: |> Output.urgent_message o Pretty.string_of bulwahn@44750: o pretty_counterex (Toplevel.context_of state) false; haftmann@28309: wenzelm@41765: val parse_arg = wenzelm@41765: Parse.name -- wenzelm@41765: (Scan.optional (Parse.$$$ "=" |-- wenzelm@41765: (((Parse.name || Parse.float_number) >> single) || wenzelm@41765: (Parse.$$$ "[" |-- Parse.list1 Parse.name --| Parse.$$$ "]"))) ["true"]); haftmann@28309: wenzelm@41765: val parse_args = wenzelm@41765: Parse.$$$ "[" |-- Parse.list1 parse_arg --| Parse.$$$ "]" || Scan.succeed []; haftmann@28336: wenzelm@36970: val _ = blanchet@43861: Outer_Syntax.command quickcheck_paramsN "set parameters for random testing" Keyword.thy_decl wenzelm@36970: (parse_args >> (fn args => Toplevel.theory (quickcheck_params_cmd args))); haftmann@28309: wenzelm@36970: val _ = blanchet@43861: Outer_Syntax.improper_command quickcheckN "try to find counterexample for subgoal" Keyword.diag wenzelm@36970: (parse_args -- Scan.optional Parse.nat 1 wenzelm@36970: >> (fn (args, i) => Toplevel.no_timing o Toplevel.keep (quickcheck_cmd args i))); haftmann@28309: blanchet@43861: (* automatic testing *) blanchet@43861: blanchet@43861: fun try_quickcheck auto state = blanchet@43861: let blanchet@43861: val ctxt = Proof.context_of state; blanchet@43861: val i = 1; blanchet@43861: val res = blanchet@43861: state blanchet@43861: |> Proof.map_context (Config.put report false #> Config.put quiet true) blanchet@43861: |> try (test_goal (false, false) ([], []) i); blanchet@43861: in blanchet@43861: case res of blanchet@43861: NONE => (unknownN, state) blanchet@43861: | SOME results => blanchet@43861: let bulwahn@44750: val msg = pretty_counterex ctxt auto (Option.map (the o get_first response_of) results) blanchet@43861: in bulwahn@44750: if is_some results then blanchet@43861: (genuineN, blanchet@43870: state blanchet@43870: |> (if auto then blanchet@43870: Proof.goal_message (K (Pretty.chunks [Pretty.str "", blanchet@43870: Pretty.mark Markup.hilite msg])) blanchet@43870: else blanchet@43870: tap (fn _ => Output.urgent_message (Pretty.string_of msg)))) blanchet@43861: else blanchet@43861: (noneN, state) blanchet@43861: end blanchet@43861: end blanchet@43861: |> `(fn (outcome_code, _) => outcome_code = genuineN) blanchet@43861: blanchet@43870: val setup = Try.register_tool (quickcheckN, (20, auto, try_quickcheck)) blanchet@43861: haftmann@28309: end; haftmann@28256: haftmann@28315: val auto_quickcheck = Quickcheck.auto;