| author | wenzelm | 
| Fri, 14 Jul 2023 14:25:53 +0200 | |
| changeset 78342 | ef1664be143d | 
| parent 74870 | d54b3c96ee50 | 
| child 78705 | fde0b195cb7d | 
| permissions | -rw-r--r-- | 
| 30824 | 1  | 
(* Title: Tools/quickcheck.ML  | 
| 
40246
 
c03fc7d3fa97
changed global fixed timeout to a configurable timeout for quickcheck; test parameters in quickcheck are now fully passed around with the context
 
bulwahn 
parents: 
40225 
diff
changeset
 | 
2  | 
Author: Stefan Berghofer, Florian Haftmann, Lukas Bulwahn, TU Muenchen  | 
| 28256 | 3  | 
|
4  | 
Generic counterexample search engine.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature QUICKCHECK =  | 
|
8  | 
sig  | 
|
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
9  | 
val quickcheckN: string  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
10  | 
val genuineN: string  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
11  | 
val noneN: string  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
12  | 
val unknownN: string  | 
| 51302 | 13  | 
(*configuration*)  | 
| 
43882
 
05d5696f177f
renaming quickcheck_tester to quickcheck_batch_tester; tuned
 
bulwahn 
parents: 
43881 
diff
changeset
 | 
14  | 
val batch_tester : string Config.T  | 
| 
40644
 
0850a2a16dce
changed old-style quickcheck configurations to new Config.T configurations
 
bulwahn 
parents: 
40643 
diff
changeset
 | 
15  | 
val size : int Config.T  | 
| 
 
0850a2a16dce
changed old-style quickcheck configurations to new Config.T configurations
 
bulwahn 
parents: 
40643 
diff
changeset
 | 
16  | 
val iterations : int Config.T  | 
| 45213 | 17  | 
val depth : int Config.T  | 
| 
40644
 
0850a2a16dce
changed old-style quickcheck configurations to new Config.T configurations
 
bulwahn 
parents: 
40643 
diff
changeset
 | 
18  | 
val no_assms : bool Config.T  | 
| 
 
0850a2a16dce
changed old-style quickcheck configurations to new Config.T configurations
 
bulwahn 
parents: 
40643 
diff
changeset
 | 
19  | 
val report : bool Config.T  | 
| 46565 | 20  | 
val timeout : real Config.T  | 
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
21  | 
val timing : bool Config.T  | 
| 
45757
 
e32dd098f57a
renaming potential flag to genuine_only flag with an inverse semantics
 
bulwahn 
parents: 
45755 
diff
changeset
 | 
22  | 
val genuine_only : bool Config.T  | 
| 51302 | 23  | 
val abort_potential : bool Config.T  | 
| 
40644
 
0850a2a16dce
changed old-style quickcheck configurations to new Config.T configurations
 
bulwahn 
parents: 
40643 
diff
changeset
 | 
24  | 
val quiet : bool Config.T  | 
| 45764 | 25  | 
val verbose : bool Config.T  | 
| 46565 | 26  | 
val use_subtype : bool Config.T  | 
27  | 
val allow_function_inversion : bool Config.T  | 
|
| 
40648
 
1598ec648b0d
splitting test_goal function in two functions; exporting new configurations in quickcheck; iterations depend on generator_name in quickcheck
 
bulwahn 
parents: 
40647 
diff
changeset
 | 
28  | 
val finite_types : bool Config.T  | 
| 
 
1598ec648b0d
splitting test_goal function in two functions; exporting new configurations in quickcheck; iterations depend on generator_name in quickcheck
 
bulwahn 
parents: 
40647 
diff
changeset
 | 
29  | 
val finite_type_size : int Config.T  | 
| 46863 | 30  | 
val tag : string Config.T  | 
| 
47348
 
9a82999ebbd6
added option quickcheck_locale to allow different behaviours for handling locales in Quickcheck;
 
bulwahn 
parents: 
46961 
diff
changeset
 | 
31  | 
val locale : string Config.T  | 
| 
43912
 
13e6a4e70219
exporting function in quickcheck; adapting mutabelle script
 
bulwahn 
parents: 
43884 
diff
changeset
 | 
32  | 
val set_active_testers: string list -> Context.generic -> Context.generic  | 
| 41517 | 33  | 
datatype expectation = No_Expectation | No_Counterexample | Counterexample;  | 
| 
40644
 
0850a2a16dce
changed old-style quickcheck configurations to new Config.T configurations
 
bulwahn 
parents: 
40643 
diff
changeset
 | 
34  | 
  datatype test_params = Test_Params of {default_type: typ list, expect : expectation};
 | 
| 
40246
 
c03fc7d3fa97
changed global fixed timeout to a configurable timeout for quickcheck; test parameters in quickcheck are now fully passed around with the context
 
bulwahn 
parents: 
40225 
diff
changeset
 | 
35  | 
val test_params_of : Proof.context -> test_params  | 
| 
40644
 
0850a2a16dce
changed old-style quickcheck configurations to new Config.T configurations
 
bulwahn 
parents: 
40643 
diff
changeset
 | 
36  | 
val map_test_params : (typ list * expectation -> typ list * expectation)  | 
| 
40246
 
c03fc7d3fa97
changed global fixed timeout to a configurable timeout for quickcheck; test parameters in quickcheck are now fully passed around with the context
 
bulwahn 
parents: 
40225 
diff
changeset
 | 
37  | 
-> Context.generic -> Context.generic  | 
| 
45159
 
3f1d1ce024cb
moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
 
bulwahn 
parents: 
43915 
diff
changeset
 | 
38  | 
val default_type : Proof.context -> typ list  | 
| 
42089
 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 
bulwahn 
parents: 
42088 
diff
changeset
 | 
39  | 
datatype report = Report of  | 
| 
 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 
bulwahn 
parents: 
42088 
diff
changeset
 | 
40  | 
    { iterations : int, raised_match_errors : int,
 | 
| 
 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 
bulwahn 
parents: 
42088 
diff
changeset
 | 
41  | 
satisfied_assms : int list, positive_concl_tests : int }  | 
| 51302 | 42  | 
(*quickcheck's result*)  | 
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
43  | 
datatype result =  | 
| 
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
44  | 
Result of  | 
| 
45727
 
5e46c225370e
extending quickcheck's result by the genuine flag
 
bulwahn 
parents: 
45682 
diff
changeset
 | 
45  | 
     {counterexample : (bool * (string * term) list) option,
 | 
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
46  | 
evaluation_terms : (term * term) list option,  | 
| 
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
47  | 
timings : (string * int) list,  | 
| 
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
48  | 
reports : (int * report) list}  | 
| 
43314
 
a9090cabca14
adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
 
bulwahn 
parents: 
43147 
diff
changeset
 | 
49  | 
val empty_result : result  | 
| 
45159
 
3f1d1ce024cb
moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
 
bulwahn 
parents: 
43915 
diff
changeset
 | 
50  | 
val found_counterexample : result -> bool  | 
| 43585 | 51  | 
val add_timing : (string * int) -> result Unsynchronized.ref -> unit  | 
| 51302 | 52  | 
val add_response : string list -> term list -> (bool * term list) option ->  | 
53  | 
result Unsynchronized.ref -> unit  | 
|
| 
45159
 
3f1d1ce024cb
moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
 
bulwahn 
parents: 
43915 
diff
changeset
 | 
54  | 
val add_report : int -> report option -> result Unsynchronized.ref -> unit  | 
| 
45727
 
5e46c225370e
extending quickcheck's result by the genuine flag
 
bulwahn 
parents: 
45682 
diff
changeset
 | 
55  | 
val counterexample_of : result -> (bool * (string * term) list) option  | 
| 
42089
 
904897d0c5bd
adapting mutabelle; exporting more Quickcheck functions
 
bulwahn 
parents: 
42088 
diff
changeset
 | 
56  | 
val timings_of : result -> (string * int) list  | 
| 51302 | 57  | 
(*registering testers & generators*)  | 
| 
43878
 
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
 
bulwahn 
parents: 
43877 
diff
changeset
 | 
58  | 
type tester =  | 
| 
45419
 
10ba32c347b0
quickcheck fails with code generator errors only if one tester is invoked
 
bulwahn 
parents: 
45418 
diff
changeset
 | 
59  | 
Proof.context -> bool -> (string * typ) list -> (term * term list) list -> result list  | 
| 
43878
 
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
 
bulwahn 
parents: 
43877 
diff
changeset
 | 
60  | 
val add_tester : string * (bool Config.T * tester) -> Context.generic -> Context.generic  | 
| 
 
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
 
bulwahn 
parents: 
43877 
diff
changeset
 | 
61  | 
val add_batch_generator :  | 
| 
43112
 
3117573292b8
adding registration of testers in Quickcheck for its use in Quickcheck_Narrowing
 
bulwahn 
parents: 
43029 
diff
changeset
 | 
62  | 
string * (Proof.context -> term list -> (int -> term list option) list)  | 
| 
 
3117573292b8
adding registration of testers in Quickcheck for its use in Quickcheck_Narrowing
 
bulwahn 
parents: 
43029 
diff
changeset
 | 
63  | 
-> Context.generic -> Context.generic  | 
| 
43878
 
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
 
bulwahn 
parents: 
43877 
diff
changeset
 | 
64  | 
val add_batch_validator :  | 
| 
43112
 
3117573292b8
adding registration of testers in Quickcheck for its use in Quickcheck_Narrowing
 
bulwahn 
parents: 
43029 
diff
changeset
 | 
65  | 
string * (Proof.context -> term list -> (int -> bool) list)  | 
| 
 
3117573292b8
adding registration of testers in Quickcheck for its use in Quickcheck_Narrowing
 
bulwahn 
parents: 
43029 
diff
changeset
 | 
66  | 
-> Context.generic -> Context.generic  | 
| 51302 | 67  | 
(*basic operations*)  | 
| 
45159
 
3f1d1ce024cb
moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
 
bulwahn 
parents: 
43915 
diff
changeset
 | 
68  | 
val message : Proof.context -> string -> unit  | 
| 
45765
 
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
 
bulwahn 
parents: 
45764 
diff
changeset
 | 
69  | 
val verbose_message : Proof.context -> string -> unit  | 
| 
45159
 
3f1d1ce024cb
moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
 
bulwahn 
parents: 
43915 
diff
changeset
 | 
70  | 
val limit : Time.time -> (bool * bool) -> (unit -> 'a) -> (unit -> 'a) -> unit -> 'a  | 
| 
45755
 
b27a06dfb2ef
outputing the potentially spurious counterexample and continue search
 
bulwahn 
parents: 
45730 
diff
changeset
 | 
71  | 
val pretty_counterex : Proof.context -> bool ->  | 
| 
 
b27a06dfb2ef
outputing the potentially spurious counterexample and continue search
 
bulwahn 
parents: 
45730 
diff
changeset
 | 
72  | 
((bool * (string * term) list) * (term * term) list) option -> Pretty.T  | 
| 51302 | 73  | 
(*testing terms and proof states*)  | 
| 
45159
 
3f1d1ce024cb
moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
 
bulwahn 
parents: 
43915 
diff
changeset
 | 
74  | 
val mk_batch_validator : Proof.context -> term list -> (int -> bool) list option  | 
| 
 
3f1d1ce024cb
moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
 
bulwahn 
parents: 
43915 
diff
changeset
 | 
75  | 
val mk_batch_tester : Proof.context -> term list -> (int -> term list option) list option  | 
| 
 
3f1d1ce024cb
moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
 
bulwahn 
parents: 
43915 
diff
changeset
 | 
76  | 
val active_testers : Proof.context -> tester list  | 
| 51302 | 77  | 
val test_terms : Proof.context -> bool * bool -> (string * typ) list ->  | 
78  | 
(term * term list) list -> result list option  | 
|
79  | 
val quickcheck: (string * string list) list -> int -> Proof.state ->  | 
|
80  | 
(bool * (string * term) list) option  | 
|
| 28256 | 81  | 
end;  | 
82  | 
||
83  | 
structure Quickcheck : QUICKCHECK =  | 
|
84  | 
struct  | 
|
85  | 
||
| 55627 | 86  | 
val quickcheckN = "quickcheck";  | 
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
87  | 
|
| 55627 | 88  | 
val genuineN = "genuine";  | 
89  | 
val noneN = "none";  | 
|
90  | 
val unknownN = "unknown";  | 
|
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
91  | 
|
| 51302 | 92  | 
|
| 
35378
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35324 
diff
changeset
 | 
93  | 
(* quickcheck report *)  | 
| 
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35324 
diff
changeset
 | 
94  | 
|
| 
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35324 
diff
changeset
 | 
95  | 
datatype report = Report of  | 
| 55627 | 96  | 
 {iterations : int,
 | 
97  | 
raised_match_errors : int,  | 
|
98  | 
satisfied_assms : int list,  | 
|
99  | 
positive_concl_tests : int};  | 
|
| 
35378
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35324 
diff
changeset
 | 
100  | 
|
| 51302 | 101  | 
|
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
102  | 
(* Quickcheck Result *)  | 
| 
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
103  | 
|
| 
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
104  | 
datatype result = Result of  | 
| 55627 | 105  | 
 {counterexample : (bool * (string * term) list) option,
 | 
106  | 
evaluation_terms : (term * term) list option,  | 
|
107  | 
timings : (string * int) list,  | 
|
108  | 
reports : (int * report) list};  | 
|
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
109  | 
|
| 
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
110  | 
val empty_result =  | 
| 55627 | 111  | 
  Result {counterexample = NONE, evaluation_terms = NONE, timings = [], reports = []};
 | 
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
112  | 
|
| 55627 | 113  | 
fun counterexample_of (Result r) = #counterexample r;  | 
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
114  | 
|
| 55627 | 115  | 
fun found_counterexample (Result r) = is_some (#counterexample r);  | 
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
116  | 
|
| 55627 | 117  | 
fun response_of (Result r) =  | 
118  | 
(case (#counterexample r, #evaluation_terms r) of  | 
|
| 
45730
 
6bd0acefaedb
outputing if counterexample is potentially spurious or not
 
bulwahn 
parents: 
45727 
diff
changeset
 | 
119  | 
(SOME ts, SOME evals) => SOME (ts, evals)  | 
| 55627 | 120  | 
| (NONE, NONE) => NONE);  | 
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
121  | 
|
| 55627 | 122  | 
fun timings_of (Result r) = #timings r;  | 
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
123  | 
|
| 
45727
 
5e46c225370e
extending quickcheck's result by the genuine flag
 
bulwahn 
parents: 
45682 
diff
changeset
 | 
124  | 
fun set_response names eval_terms (SOME (genuine, ts)) (Result r) =  | 
| 55627 | 125  | 
let  | 
126  | 
val (ts1, ts2) = chop (length names) ts  | 
|
127  | 
val (eval_terms', _) = chop (length ts2) eval_terms  | 
|
128  | 
in  | 
|
129  | 
        Result {counterexample = SOME (genuine, (names ~~ ts1)),
 | 
|
130  | 
evaluation_terms = SOME (eval_terms' ~~ ts2),  | 
|
131  | 
timings = #timings r, reports = #reports r}  | 
|
132  | 
end  | 
|
133  | 
| set_response _ _ NONE result = result;  | 
|
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
134  | 
|
| 43585 | 135  | 
|
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
136  | 
fun cons_timing timing (Result r) =  | 
| 
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
137  | 
  Result {counterexample = #counterexample r, evaluation_terms = #evaluation_terms r,
 | 
| 55627 | 138  | 
timings = cons timing (#timings r), reports = #reports r};  | 
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
139  | 
|
| 
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
140  | 
fun cons_report size (SOME report) (Result r) =  | 
| 55627 | 141  | 
      Result {counterexample = #counterexample r, evaluation_terms = #evaluation_terms r,
 | 
142  | 
timings = #timings r, reports = cons (size, report) (#reports r)}  | 
|
143  | 
| cons_report _ NONE result = result;  | 
|
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
144  | 
|
| 
42198
 
ded5ba6b7bac
use Unsynchronized.change convenience, which also emphasizes the raw access to these references (which happen to be local here);
 
wenzelm 
parents: 
42194 
diff
changeset
 | 
145  | 
fun add_timing timing result_ref =  | 
| 55627 | 146  | 
Unsynchronized.change result_ref (cons_timing timing);  | 
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
147  | 
|
| 
42198
 
ded5ba6b7bac
use Unsynchronized.change convenience, which also emphasizes the raw access to these references (which happen to be local here);
 
wenzelm 
parents: 
42194 
diff
changeset
 | 
148  | 
fun add_report size report result_ref =  | 
| 55627 | 149  | 
Unsynchronized.change result_ref (cons_report size report);  | 
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
150  | 
|
| 
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
151  | 
fun add_response names eval_terms response result_ref =  | 
| 55627 | 152  | 
Unsynchronized.change result_ref (set_response names eval_terms response);  | 
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
153  | 
|
| 51302 | 154  | 
|
| 
37929
 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 
bulwahn 
parents: 
37913 
diff
changeset
 | 
155  | 
(* expectation *)  | 
| 
 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 
bulwahn 
parents: 
37913 
diff
changeset
 | 
156  | 
|
| 41517 | 157  | 
datatype expectation = No_Expectation | No_Counterexample | Counterexample;  | 
| 
37929
 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 
bulwahn 
parents: 
37913 
diff
changeset
 | 
158  | 
|
| 
 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 
bulwahn 
parents: 
37913 
diff
changeset
 | 
159  | 
fun merge_expectation (expect1, expect2) =  | 
| 55627 | 160  | 
if expect1 = expect2 then expect1 else No_Expectation;  | 
| 
37929
 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 
bulwahn 
parents: 
37913 
diff
changeset
 | 
161  | 
|
| 51302 | 162  | 
(*quickcheck configuration -- default parameters, test generators*)  | 
| 67149 | 163  | 
val batch_tester = Attrib.setup_config_string \<^binding>\<open>quickcheck_batch_tester\<close> (K "");  | 
164  | 
val size = Attrib.setup_config_int \<^binding>\<open>quickcheck_size\<close> (K 10);  | 
|
165  | 
val iterations = Attrib.setup_config_int \<^binding>\<open>quickcheck_iterations\<close> (K 100);  | 
|
166  | 
val depth = Attrib.setup_config_int \<^binding>\<open>quickcheck_depth\<close> (K 10);  | 
|
| 45213 | 167  | 
|
| 67149 | 168  | 
val no_assms = Attrib.setup_config_bool \<^binding>\<open>quickcheck_no_assms\<close> (K false);  | 
169  | 
val locale = Attrib.setup_config_string \<^binding>\<open>quickcheck_locale\<close> (K "interpret expand");  | 
|
170  | 
val report = Attrib.setup_config_bool \<^binding>\<open>quickcheck_report\<close> (K true);  | 
|
171  | 
val timing = Attrib.setup_config_bool \<^binding>\<open>quickcheck_timing\<close> (K false);  | 
|
172  | 
val timeout = Attrib.setup_config_real \<^binding>\<open>quickcheck_timeout\<close> (K 30.0);  | 
|
| 46565 | 173  | 
|
| 67149 | 174  | 
val genuine_only = Attrib.setup_config_bool \<^binding>\<open>quickcheck_genuine_only\<close> (K false);  | 
175  | 
val abort_potential = Attrib.setup_config_bool \<^binding>\<open>quickcheck_abort_potential\<close> (K false);  | 
|
| 46565 | 176  | 
|
| 67149 | 177  | 
val quiet = Attrib.setup_config_bool \<^binding>\<open>quickcheck_quiet\<close> (K false);  | 
178  | 
val verbose = Attrib.setup_config_bool \<^binding>\<open>quickcheck_verbose\<close> (K false);  | 
|
179  | 
val tag = Attrib.setup_config_string \<^binding>\<open>quickcheck_tag\<close> (K "");  | 
|
| 46565 | 180  | 
|
| 67149 | 181  | 
val use_subtype = Attrib.setup_config_bool \<^binding>\<open>quickcheck_use_subtype\<close> (K false);  | 
| 46565 | 182  | 
|
| 
45449
 
eeffd04cd899
adding option allow_function_inversion to quickcheck options
 
bulwahn 
parents: 
45419 
diff
changeset
 | 
183  | 
val allow_function_inversion =  | 
| 67149 | 184  | 
Attrib.setup_config_bool \<^binding>\<open>quickcheck_allow_function_inversion\<close> (K false);  | 
185  | 
val finite_types = Attrib.setup_config_bool \<^binding>\<open>quickcheck_finite_types\<close> (K true);  | 
|
186  | 
val finite_type_size = Attrib.setup_config_int \<^binding>\<open>quickcheck_finite_type_size\<close> (K 3);  | 
|
| 40646 | 187  | 
|
| 
40644
 
0850a2a16dce
changed old-style quickcheck configurations to new Config.T configurations
 
bulwahn 
parents: 
40643 
diff
changeset
 | 
188  | 
datatype test_params = Test_Params of  | 
| 
 
0850a2a16dce
changed old-style quickcheck configurations to new Config.T configurations
 
bulwahn 
parents: 
40643 
diff
changeset
 | 
189  | 
  {default_type: typ list, expect : expectation};
 | 
| 
38759
 
37a9092de102
simplification/standardization of some theory data;
 
wenzelm 
parents: 
38390 
diff
changeset
 | 
190  | 
|
| 
40644
 
0850a2a16dce
changed old-style quickcheck configurations to new Config.T configurations
 
bulwahn 
parents: 
40643 
diff
changeset
 | 
191  | 
fun dest_test_params (Test_Params {default_type, expect}) = (default_type, expect);
 | 
| 
 
0850a2a16dce
changed old-style quickcheck configurations to new Config.T configurations
 
bulwahn 
parents: 
40643 
diff
changeset
 | 
192  | 
|
| 41517 | 193  | 
fun make_test_params (default_type, expect) =  | 
194  | 
  Test_Params {default_type = default_type, expect = expect};
 | 
|
| 
40644
 
0850a2a16dce
changed old-style quickcheck configurations to new Config.T configurations
 
bulwahn 
parents: 
40643 
diff
changeset
 | 
195  | 
|
| 41517 | 196  | 
fun map_test_params' f (Test_Params {default_type, expect}) =
 | 
197  | 
make_test_params (f (default_type, expect));  | 
|
| 
38759
 
37a9092de102
simplification/standardization of some theory data;
 
wenzelm 
parents: 
38390 
diff
changeset
 | 
198  | 
|
| 
 
37a9092de102
simplification/standardization of some theory data;
 
wenzelm 
parents: 
38390 
diff
changeset
 | 
199  | 
fun merge_test_params  | 
| 
41472
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
41086 
diff
changeset
 | 
200  | 
  (Test_Params {default_type = default_type1, expect = expect1},
 | 
| 
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
41086 
diff
changeset
 | 
201  | 
    Test_Params {default_type = default_type2, expect = expect2}) =
 | 
| 
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
41086 
diff
changeset
 | 
202  | 
make_test_params  | 
| 
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
41086 
diff
changeset
 | 
203  | 
(merge (op =) (default_type1, default_type2), merge_expectation (expect1, expect2));  | 
| 28309 | 204  | 
|
| 
43878
 
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
 
bulwahn 
parents: 
43877 
diff
changeset
 | 
205  | 
type tester =  | 
| 51302 | 206  | 
Proof.context -> bool -> (string * typ) list -> (term * term list) list -> result list;  | 
| 
43878
 
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
 
bulwahn 
parents: 
43877 
diff
changeset
 | 
207  | 
|
| 
39252
 
8f176e575a49
changing the container for the quickcheck options to a generic data
 
bulwahn 
parents: 
39138 
diff
changeset
 | 
208  | 
structure Data = Generic_Data  | 
| 33522 | 209  | 
(  | 
| 
38759
 
37a9092de102
simplification/standardization of some theory data;
 
wenzelm 
parents: 
38390 
diff
changeset
 | 
210  | 
type T =  | 
| 59436 | 211  | 
(string * (bool Config.T * tester)) list *  | 
212  | 
(string * (Proof.context -> term list -> (int -> term list option) list)) list *  | 
|
213  | 
(string * (Proof.context -> term list -> (int -> bool) list)) list *  | 
|
214  | 
test_params;  | 
|
215  | 
  val empty = ([], [], [], Test_Params {default_type = [], expect = No_Expectation});
 | 
|
| 55627 | 216  | 
fun merge  | 
| 59436 | 217  | 
((testers1, batch_generators1, batch_validators1, params1),  | 
218  | 
(testers2, batch_generators2, batch_validators2, params2)) : T =  | 
|
219  | 
(AList.merge (op =) (K true) (testers1, testers2),  | 
|
220  | 
AList.merge (op =) (K true) (batch_generators1, batch_generators2),  | 
|
221  | 
AList.merge (op =) (K true) (batch_validators1, batch_validators2),  | 
|
222  | 
merge_test_params (params1, params2));  | 
|
| 33522 | 223  | 
);  | 
| 28256 | 224  | 
|
| 59436 | 225  | 
val test_params_of = #4 o Data.get o Context.Proof;  | 
| 55627 | 226  | 
val default_type = fst o dest_test_params o test_params_of;  | 
227  | 
val expect = snd o dest_test_params o test_params_of;  | 
|
| 59436 | 228  | 
val map_test_params = Data.map o @{apply 4(4)} o map_test_params';
 | 
| 
40246
 
c03fc7d3fa97
changed global fixed timeout to a configurable timeout for quickcheck; test parameters in quickcheck are now fully passed around with the context
 
bulwahn 
parents: 
40225 
diff
changeset
 | 
229  | 
|
| 59436 | 230  | 
val add_tester = Data.map o @{apply 4(1)} o AList.update (op =);
 | 
231  | 
val add_batch_generator = Data.map o @{apply 4(2)} o AList.update (op =);
 | 
|
232  | 
val add_batch_validator = Data.map o @{apply 4(3)} o AList.update (op =);
 | 
|
| 28309 | 233  | 
|
| 
43881
 
cabe74eab19a
changing parser in quickcheck to activate and deactivate the testers
 
bulwahn 
parents: 
43879 
diff
changeset
 | 
234  | 
fun active_testers ctxt =  | 
| 
 
cabe74eab19a
changing parser in quickcheck to activate and deactivate the testers
 
bulwahn 
parents: 
43879 
diff
changeset
 | 
235  | 
let  | 
| 59436 | 236  | 
val testers = map snd (#1 (Data.get (Context.Proof ctxt)));  | 
| 
43881
 
cabe74eab19a
changing parser in quickcheck to activate and deactivate the testers
 
bulwahn 
parents: 
43879 
diff
changeset
 | 
237  | 
in  | 
| 
 
cabe74eab19a
changing parser in quickcheck to activate and deactivate the testers
 
bulwahn 
parents: 
43879 
diff
changeset
 | 
238  | 
map snd (filter (fn (active, _) => Config.get ctxt active) testers)  | 
| 55627 | 239  | 
end;  | 
| 51302 | 240  | 
|
| 55629 | 241  | 
fun set_active_testers [] context = context  | 
242  | 
| set_active_testers testers context =  | 
|
| 51302 | 243  | 
let  | 
| 59436 | 244  | 
val registered_testers = #1 (Data.get context);  | 
| 51302 | 245  | 
in  | 
246  | 
fold (fn (name, (config, _)) => Config.put_generic config (member (op =) testers name))  | 
|
| 55629 | 247  | 
registered_testers context  | 
| 51302 | 248  | 
end;  | 
249  | 
||
250  | 
||
| 28315 | 251  | 
(* generating tests *)  | 
252  | 
||
| 
41862
 
a38536bf2736
adding function Quickcheck.test_terms to provide checking a batch of terms
 
bulwahn 
parents: 
41754 
diff
changeset
 | 
253  | 
fun gen_mk_tester lookup ctxt v =  | 
| 28309 | 254  | 
let  | 
| 
43882
 
05d5696f177f
renaming quickcheck_tester to quickcheck_batch_tester; tuned
 
bulwahn 
parents: 
43881 
diff
changeset
 | 
255  | 
val name = Config.get ctxt batch_tester  | 
| 51302 | 256  | 
val tester =  | 
257  | 
(case lookup ctxt name of  | 
|
258  | 
        NONE => error ("No such quickcheck batch-tester: " ^ name)
 | 
|
259  | 
| SOME tester => tester ctxt);  | 
|
| 
40235
 
87998864284e
use Exn.interruptible_capture to keep user-code interruptible (Exn.capture not immediately followed by Exn.release here);
 
wenzelm 
parents: 
40225 
diff
changeset
 | 
260  | 
in  | 
| 
40909
 
e006d1e06920
renamed parameter from generator to tester; quickcheck only applies one tester on invocation
 
bulwahn 
parents: 
40908 
diff
changeset
 | 
261  | 
if Config.get ctxt quiet then  | 
| 
41862
 
a38536bf2736
adding function Quickcheck.test_terms to provide checking a batch of terms
 
bulwahn 
parents: 
41754 
diff
changeset
 | 
262  | 
try tester v  | 
| 
40909
 
e006d1e06920
renamed parameter from generator to tester; quickcheck only applies one tester on invocation
 
bulwahn 
parents: 
40908 
diff
changeset
 | 
263  | 
else  | 
| 
43761
 
e72ba84ae58f
tuned signature -- corresponding to Scala version;
 
wenzelm 
parents: 
43585 
diff
changeset
 | 
264  | 
let (* FIXME !?!? *)  | 
| 
41862
 
a38536bf2736
adding function Quickcheck.test_terms to provide checking a batch of terms
 
bulwahn 
parents: 
41754 
diff
changeset
 | 
265  | 
val tester = Exn.interruptible_capture tester v  | 
| 51302 | 266  | 
in  | 
267  | 
(case Exn.get_res tester of  | 
|
| 
40909
 
e006d1e06920
renamed parameter from generator to tester; quickcheck only applies one tester on invocation
 
bulwahn 
parents: 
40908 
diff
changeset
 | 
268  | 
NONE => SOME (Exn.release tester)  | 
| 51302 | 269  | 
| SOME tester => SOME tester)  | 
| 
40909
 
e006d1e06920
renamed parameter from generator to tester; quickcheck only applies one tester on invocation
 
bulwahn 
parents: 
40908 
diff
changeset
 | 
270  | 
end  | 
| 51302 | 271  | 
end;  | 
| 
43882
 
05d5696f177f
renaming quickcheck_tester to quickcheck_batch_tester; tuned
 
bulwahn 
parents: 
43881 
diff
changeset
 | 
272  | 
|
| 59436 | 273  | 
val mk_batch_tester = gen_mk_tester (AList.lookup (op =) o #2 o Data.get o Context.Proof);  | 
274  | 
val mk_batch_validator = gen_mk_tester (AList.lookup (op =) o #3 o Data.get o Context.Proof);  | 
|
| 51302 | 275  | 
|
276  | 
||
| 28315 | 277  | 
(* testing propositions *)  | 
278  | 
||
| 43876 | 279  | 
type compile_generator =  | 
| 51302 | 280  | 
Proof.context -> (term * term list) list -> int list -> term list option * report option;  | 
| 43876 | 281  | 
|
| 43585 | 282  | 
fun limit timeout (limit_time, is_interactive) f exc () =  | 
| 
41754
 
aa94a003dcdf
quickcheck can be invoked with its internal timelimit or without
 
bulwahn 
parents: 
41753 
diff
changeset
 | 
283  | 
if limit_time then  | 
| 
74870
 
d54b3c96ee50
more robust physical timeout (despite 1bea05713dde), especially relevant for quickcheck where large unary numerals may cause excessive heap allocations and resulting GC is better included in the timing;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
284  | 
Timeout.apply_physical timeout f ()  | 
| 62519 | 285  | 
handle timeout_exn as Timeout.TIMEOUT _ =>  | 
286  | 
if is_interactive then exc () else Exn.reraise timeout_exn  | 
|
| 55627 | 287  | 
else f ();  | 
| 51302 | 288  | 
|
| 58843 | 289  | 
fun message ctxt s = if Config.get ctxt quiet then () else writeln s;  | 
| 
41753
 
dbd00d8a4784
quickcheck invokes TimeLimit.timeLimit only in one separate function
 
bulwahn 
parents: 
41735 
diff
changeset
 | 
290  | 
|
| 51302 | 291  | 
fun verbose_message ctxt s =  | 
292  | 
if not (Config.get ctxt quiet) andalso Config.get ctxt verbose  | 
|
| 58843 | 293  | 
then writeln s else ();  | 
| 
45765
 
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
 
bulwahn 
parents: 
45764 
diff
changeset
 | 
294  | 
|
| 
62983
 
ba9072b303a2
avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
 
wenzelm 
parents: 
62982 
diff
changeset
 | 
295  | 
fun test_terms ctxt0 (limit_time, is_interactive) insts goals =  | 
| 
 
ba9072b303a2
avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
 
wenzelm 
parents: 
62982 
diff
changeset
 | 
296  | 
let val ctxt = Simplifier_Trace.disable ctxt0 in  | 
| 
 
ba9072b303a2
avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
 
wenzelm 
parents: 
62982 
diff
changeset
 | 
297  | 
(case active_testers ctxt of  | 
| 
 
ba9072b303a2
avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
 
wenzelm 
parents: 
62982 
diff
changeset
 | 
298  | 
[] => error "No active testers for quickcheck"  | 
| 
 
ba9072b303a2
avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
 
wenzelm 
parents: 
62982 
diff
changeset
 | 
299  | 
| testers =>  | 
| 
 
ba9072b303a2
avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
 
wenzelm 
parents: 
62982 
diff
changeset
 | 
300  | 
limit (seconds (Config.get ctxt timeout)) (limit_time, is_interactive)  | 
| 
 
ba9072b303a2
avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
 
wenzelm 
parents: 
62982 
diff
changeset
 | 
301  | 
(fn () =>  | 
| 
 
ba9072b303a2
avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
 
wenzelm 
parents: 
62982 
diff
changeset
 | 
302  | 
Par_List.get_some (fn tester =>  | 
| 
 
ba9072b303a2
avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
 
wenzelm 
parents: 
62982 
diff
changeset
 | 
303  | 
tester ctxt (length testers > 1) insts goals |>  | 
| 
 
ba9072b303a2
avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
 
wenzelm 
parents: 
62982 
diff
changeset
 | 
304  | 
(fn result => if exists found_counterexample result then SOME result else NONE))  | 
| 
 
ba9072b303a2
avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
 
wenzelm 
parents: 
62982 
diff
changeset
 | 
305  | 
testers)  | 
| 
 
ba9072b303a2
avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
 
wenzelm 
parents: 
62982 
diff
changeset
 | 
306  | 
(fn () => (message ctxt "Quickcheck ran out of time"; NONE)) ())  | 
| 
 
ba9072b303a2
avoid misleading Simplifier trace in quickcheck, notably in auto quickcheck;
 
wenzelm 
parents: 
62982 
diff
changeset
 | 
307  | 
end  | 
| 
46759
 
a6ea1c68fa52
collecting all axioms in a locale context in quickcheck;
 
bulwahn 
parents: 
46565 
diff
changeset
 | 
308  | 
|
| 
 
a6ea1c68fa52
collecting all axioms in a locale context in quickcheck;
 
bulwahn 
parents: 
46565 
diff
changeset
 | 
309  | 
fun all_axioms_of ctxt t =  | 
| 
 
a6ea1c68fa52
collecting all axioms in a locale context in quickcheck;
 
bulwahn 
parents: 
46565 
diff
changeset
 | 
310  | 
let  | 
| 51302 | 311  | 
val intros = Locale.get_intros ctxt;  | 
312  | 
val unfolds = Locale.get_unfolds ctxt;  | 
|
313  | 
fun retrieve_prems thms t =  | 
|
314  | 
(case filter (fn th => Term.could_unify (Thm.concl_of th, t)) thms of  | 
|
315  | 
[] => NONE  | 
|
| 
46759
 
a6ea1c68fa52
collecting all axioms in a locale context in quickcheck;
 
bulwahn 
parents: 
46565 
diff
changeset
 | 
316  | 
| [th] =>  | 
| 51302 | 317  | 
let  | 
318  | 
val (tyenv, tenv) =  | 
|
319  | 
Pattern.match (Proof_Context.theory_of ctxt)  | 
|
320  | 
(Thm.concl_of th, t) (Vartab.empty, Vartab.empty)  | 
|
321  | 
in SOME (map (Envir.subst_term (tyenv, tenv)) (Thm.prems_of th)) end);  | 
|
| 
46759
 
a6ea1c68fa52
collecting all axioms in a locale context in quickcheck;
 
bulwahn 
parents: 
46565 
diff
changeset
 | 
322  | 
fun all t =  | 
| 51302 | 323  | 
(case retrieve_prems intros t of  | 
| 
46759
 
a6ea1c68fa52
collecting all axioms in a locale context in quickcheck;
 
bulwahn 
parents: 
46565 
diff
changeset
 | 
324  | 
NONE => retrieve_prems unfolds t  | 
| 51302 | 325  | 
| SOME ts => SOME (maps (fn t => the_default [t] (all t)) ts));  | 
| 
46759
 
a6ea1c68fa52
collecting all axioms in a locale context in quickcheck;
 
bulwahn 
parents: 
46565 
diff
changeset
 | 
326  | 
in  | 
| 
 
a6ea1c68fa52
collecting all axioms in a locale context in quickcheck;
 
bulwahn 
parents: 
46565 
diff
changeset
 | 
327  | 
all t  | 
| 51302 | 328  | 
end;  | 
| 
46759
 
a6ea1c68fa52
collecting all axioms in a locale context in quickcheck;
 
bulwahn 
parents: 
46565 
diff
changeset
 | 
329  | 
|
| 
47348
 
9a82999ebbd6
added option quickcheck_locale to allow different behaviours for handling locales in Quickcheck;
 
bulwahn 
parents: 
46961 
diff
changeset
 | 
330  | 
fun locale_config_of s =  | 
| 
 
9a82999ebbd6
added option quickcheck_locale to allow different behaviours for handling locales in Quickcheck;
 
bulwahn 
parents: 
46961 
diff
changeset
 | 
331  | 
let  | 
| 51302 | 332  | 
val cs = space_explode " " s;  | 
| 
47348
 
9a82999ebbd6
added option quickcheck_locale to allow different behaviours for handling locales in Quickcheck;
 
bulwahn 
parents: 
46961 
diff
changeset
 | 
333  | 
in  | 
| 
 
9a82999ebbd6
added option quickcheck_locale to allow different behaviours for handling locales in Quickcheck;
 
bulwahn 
parents: 
46961 
diff
changeset
 | 
334  | 
if forall (fn c => c = "expand" orelse c = "interpret") cs then cs  | 
| 55627 | 335  | 
else  | 
336  | 
     (warning ("Invalid quickcheck_locale setting: falling back to the default setting.");
 | 
|
| 
47348
 
9a82999ebbd6
added option quickcheck_locale to allow different behaviours for handling locales in Quickcheck;
 
bulwahn 
parents: 
46961 
diff
changeset
 | 
337  | 
["interpret", "expand"])  | 
| 51302 | 338  | 
end;  | 
| 
47348
 
9a82999ebbd6
added option quickcheck_locale to allow different behaviours for handling locales in Quickcheck;
 
bulwahn 
parents: 
46961 
diff
changeset
 | 
339  | 
|
| 
42026
 
da9b58f1db42
adding option of evaluating terms after invocation in quickcheck
 
bulwahn 
parents: 
42025 
diff
changeset
 | 
340  | 
fun test_goal (time_limit, is_interactive) (insts, eval_terms) i state =  | 
| 
40648
 
1598ec648b0d
splitting test_goal function in two functions; exporting new configurations in quickcheck; iterations depend on generator_name in quickcheck
 
bulwahn 
parents: 
40647 
diff
changeset
 | 
341  | 
let  | 
| 62982 | 342  | 
val ctxt = Proof.context_of state;  | 
| 
40648
 
1598ec648b0d
splitting test_goal function in two functions; exporting new configurations in quickcheck; iterations depend on generator_name in quickcheck
 
bulwahn 
parents: 
40647 
diff
changeset
 | 
343  | 
val thy = Proof.theory_of state;  | 
| 62982 | 344  | 
val opt_locale = Named_Target.bottom_locale_of ctxt;  | 
| 
73583
 
ed5226fdf89d
proper context variable handling when stripping leadings quantifiers from test goals
 
haftmann 
parents: 
67149 
diff
changeset
 | 
345  | 
fun axioms_of locale =  | 
| 
 
ed5226fdf89d
proper context variable handling when stripping leadings quantifiers from test goals
 
haftmann 
parents: 
67149 
diff
changeset
 | 
346  | 
(case fst (Locale.specification_of thy locale) of  | 
| 
 
ed5226fdf89d
proper context variable handling when stripping leadings quantifiers from test goals
 
haftmann 
parents: 
67149 
diff
changeset
 | 
347  | 
NONE => []  | 
| 
 
ed5226fdf89d
proper context variable handling when stripping leadings quantifiers from test goals
 
haftmann 
parents: 
67149 
diff
changeset
 | 
348  | 
| SOME t => the_default [] (all_axioms_of ctxt t));  | 
| 
 
ed5226fdf89d
proper context variable handling when stripping leadings quantifiers from test goals
 
haftmann 
parents: 
67149 
diff
changeset
 | 
349  | 
val config = locale_config_of (Config.get ctxt locale);  | 
| 51302 | 350  | 
val assms =  | 
| 62982 | 351  | 
if Config.get ctxt no_assms then []  | 
| 51302 | 352  | 
else  | 
| 62982 | 353  | 
(case opt_locale of  | 
354  | 
NONE => Assumption.all_assms_of ctxt  | 
|
355  | 
| SOME locale => Assumption.local_assms_of ctxt (Locale.init locale thy));  | 
|
| 
73583
 
ed5226fdf89d
proper context variable handling when stripping leadings quantifiers from test goals
 
haftmann 
parents: 
67149 
diff
changeset
 | 
356  | 
    val {goal = st, ...} = Proof.raw_goal state;
 | 
| 
 
ed5226fdf89d
proper context variable handling when stripping leadings quantifiers from test goals
 
haftmann 
parents: 
67149 
diff
changeset
 | 
357  | 
val (gi, _) = Logic.goal_params (Thm.prop_of st) i;  | 
| 
 
ed5226fdf89d
proper context variable handling when stripping leadings quantifiers from test goals
 
haftmann 
parents: 
67149 
diff
changeset
 | 
358  | 
val ((_, raw_goal), ctxt') = Variable.focus NONE gi ctxt;  | 
| 
 
ed5226fdf89d
proper context variable handling when stripping leadings quantifiers from test goals
 
haftmann 
parents: 
67149 
diff
changeset
 | 
359  | 
val proto_goal = Logic.list_implies (map Thm.term_of assms, raw_goal);  | 
| 55627 | 360  | 
val goals =  | 
| 62982 | 361  | 
(case opt_locale of  | 
| 55627 | 362  | 
NONE => [(proto_goal, eval_terms)]  | 
363  | 
| SOME locale =>  | 
|
364  | 
fold (fn c =>  | 
|
365  | 
if c = "expand" then  | 
|
366  | 
cons (Logic.list_implies (axioms_of locale, proto_goal), eval_terms)  | 
|
367  | 
else if c = "interpret" then  | 
|
368  | 
append (map (fn (_, phi) =>  | 
|
369  | 
(Morphism.term phi proto_goal, map (Morphism.term phi) eval_terms))  | 
|
370  | 
(Locale.registrations_of (Context.Theory thy) (* FIXME !? *) locale))  | 
|
371  | 
else I) config []);  | 
|
| 51302 | 372  | 
val _ =  | 
| 
73583
 
ed5226fdf89d
proper context variable handling when stripping leadings quantifiers from test goals
 
haftmann 
parents: 
67149 
diff
changeset
 | 
373  | 
verbose_message ctxt'  | 
| 51302 | 374  | 
(Pretty.string_of  | 
| 
73583
 
ed5226fdf89d
proper context variable handling when stripping leadings quantifiers from test goals
 
haftmann 
parents: 
67149 
diff
changeset
 | 
375  | 
          (Pretty.big_list ("Checking goals: ") (map (Syntax.pretty_term ctxt' o fst) goals)));
 | 
| 
40648
 
1598ec648b0d
splitting test_goal function in two functions; exporting new configurations in quickcheck; iterations depend on generator_name in quickcheck
 
bulwahn 
parents: 
40647 
diff
changeset
 | 
376  | 
in  | 
| 
73583
 
ed5226fdf89d
proper context variable handling when stripping leadings quantifiers from test goals
 
haftmann 
parents: 
67149 
diff
changeset
 | 
377  | 
test_terms ctxt' (time_limit, is_interactive) insts goals  | 
| 51302 | 378  | 
end;  | 
379  | 
||
| 
40648
 
1598ec648b0d
splitting test_goal function in two functions; exporting new configurations in quickcheck; iterations depend on generator_name in quickcheck
 
bulwahn 
parents: 
40647 
diff
changeset
 | 
380  | 
|
| 37912 | 381  | 
(* pretty printing *)  | 
| 28315 | 382  | 
|
| 59435 | 383  | 
fun tool_name auto = if auto then "Auto Quickcheck" else "Quickcheck";  | 
| 40225 | 384  | 
|
| 46863 | 385  | 
fun pretty_counterex ctxt auto NONE =  | 
| 59435 | 386  | 
Pretty.para (tool_name auto ^ " found no counterexample." ^ Config.get ctxt tag)  | 
| 
45730
 
6bd0acefaedb
outputing if counterexample is potentially spurious or not
 
bulwahn 
parents: 
45727 
diff
changeset
 | 
387  | 
| pretty_counterex ctxt auto (SOME ((genuine, cex), eval_terms)) =  | 
| 59438 | 388  | 
let  | 
389  | 
val header =  | 
|
390  | 
Pretty.para  | 
|
391  | 
(tool_name auto ^ " found a " ^  | 
|
| 60666 | 392  | 
(if genuine then "counterexample"  | 
393  | 
else "potentially spurious counterexample due to underspecified functions") ^  | 
|
394  | 
(if null cex then "." else ":") ^  | 
|
| 59438 | 395  | 
Config.get ctxt tag);  | 
| 59439 | 396  | 
fun pretty_cex (x, t) =  | 
397  | 
Pretty.block [Pretty.str (x ^ " ="), Pretty.brk 1, Syntax.pretty_term ctxt t];  | 
|
| 59438 | 398  | 
in  | 
| 59439 | 399  | 
Pretty.chunks (Pretty.block (Pretty.fbreaks (header :: map pretty_cex (rev cex))) ::  | 
| 59438 | 400  | 
(if null eval_terms then []  | 
401  | 
else  | 
|
402  | 
[Pretty.big_list "Evaluated terms:"  | 
|
403  | 
(map (fn (t, u) =>  | 
|
404  | 
Pretty.block [Syntax.pretty_term ctxt t, Pretty.str " =", Pretty.brk 1,  | 
|
405  | 
Syntax.pretty_term ctxt u]) (rev eval_terms))]))  | 
|
406  | 
end;  | 
|
| 28315 | 407  | 
|
| 51302 | 408  | 
|
| 30980 | 409  | 
(* Isar commands *)  | 
| 28315 | 410  | 
|
| 51302 | 411  | 
fun read_nat s =  | 
| 55627 | 412  | 
(case Library.read_int (Symbol.explode s) of  | 
| 51302 | 413  | 
(k, []) =>  | 
414  | 
if k >= 0 then k  | 
|
| 28336 | 415  | 
      else error ("Not a natural number: " ^ s)
 | 
| 55627 | 416  | 
  | _ => error ("Not a natural number: " ^ s));
 | 
| 
37909
 
583543ad6ad1
changed default types to a list of types; extended quickcheck parameters to be a list of values to parse a list of default types
 
bulwahn 
parents: 
36960 
diff
changeset
 | 
417  | 
|
| 
34128
 
8650a073dd9b
made Quickcheck take structured proof assumptions into account (like Refute and Nitpick) by default;
 
blanchet 
parents: 
33583 
diff
changeset
 | 
418  | 
fun read_bool "false" = false  | 
| 
 
8650a073dd9b
made Quickcheck take structured proof assumptions into account (like Refute and Nitpick) by default;
 
blanchet 
parents: 
33583 
diff
changeset
 | 
419  | 
| read_bool "true" = true  | 
| 51302 | 420  | 
  | read_bool s = error ("Not a Boolean value: " ^ s);
 | 
| 28315 | 421  | 
|
| 
40366
 
a2866dbfbe6b
changing timeout to real value; handling Interrupt and Timeout more like nitpick does
 
bulwahn 
parents: 
40253 
diff
changeset
 | 
422  | 
fun read_real s =  | 
| 55627 | 423  | 
(case Real.fromString s of  | 
| 
40366
 
a2866dbfbe6b
changing timeout to real value; handling Interrupt and Timeout more like nitpick does
 
bulwahn 
parents: 
40253 
diff
changeset
 | 
424  | 
SOME s => s  | 
| 51302 | 425  | 
  | NONE => error ("Not a real number: " ^ s));
 | 
| 
40366
 
a2866dbfbe6b
changing timeout to real value; handling Interrupt and Timeout more like nitpick does
 
bulwahn 
parents: 
40253 
diff
changeset
 | 
426  | 
|
| 
37929
 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 
bulwahn 
parents: 
37913 
diff
changeset
 | 
427  | 
fun read_expectation "no_expectation" = No_Expectation  | 
| 41517 | 428  | 
| read_expectation "no_counterexample" = No_Counterexample  | 
| 
37929
 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 
bulwahn 
parents: 
37913 
diff
changeset
 | 
429  | 
| read_expectation "counterexample" = Counterexample  | 
| 51302 | 430  | 
  | read_expectation s = error ("Not an expectation value: " ^ s);
 | 
| 
37929
 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 
bulwahn 
parents: 
37913 
diff
changeset
 | 
431  | 
|
| 59437 | 432  | 
fun valid_tester_name context name =  | 
433  | 
AList.defined (op =) (#1 (Data.get context)) name;  | 
|
| 51302 | 434  | 
|
| 59437 | 435  | 
fun parse_tester name (testers, context) =  | 
436  | 
if valid_tester_name context name then  | 
|
437  | 
(insert (op =) name testers, context)  | 
|
| 55627 | 438  | 
  else error ("Unknown tester: " ^ name);
 | 
| 
40912
 
1108529100ce
checking if parameter is name of a tester which allows e.g. quickcheck[random]
 
bulwahn 
parents: 
40911 
diff
changeset
 | 
439  | 
|
| 
43881
 
cabe74eab19a
changing parser in quickcheck to activate and deactivate the testers
 
bulwahn 
parents: 
43879 
diff
changeset
 | 
440  | 
fun parse_test_param ("tester", args) = fold parse_tester args
 | 
| 
 
cabe74eab19a
changing parser in quickcheck to activate and deactivate the testers
 
bulwahn 
parents: 
43879 
diff
changeset
 | 
441  | 
  | parse_test_param ("size", [arg]) = apsnd (Config.put_generic size (read_nat arg))
 | 
| 
 
cabe74eab19a
changing parser in quickcheck to activate and deactivate the testers
 
bulwahn 
parents: 
43879 
diff
changeset
 | 
442  | 
  | parse_test_param ("iterations", [arg]) = apsnd (Config.put_generic iterations (read_nat arg))
 | 
| 51302 | 443  | 
  | parse_test_param ("depth", [arg]) = apsnd (Config.put_generic depth (read_nat arg))
 | 
| 55627 | 444  | 
  | parse_test_param ("default_type", arg) =
 | 
| 55629 | 445  | 
(fn (testers, context) =>  | 
| 55627 | 446  | 
(testers, map_test_params  | 
| 55629 | 447  | 
(apfst (K (map (Proof_Context.read_typ (Context.proof_of context)) arg))) context))  | 
| 
43881
 
cabe74eab19a
changing parser in quickcheck to activate and deactivate the testers
 
bulwahn 
parents: 
43879 
diff
changeset
 | 
448  | 
  | parse_test_param ("no_assms", [arg]) = apsnd (Config.put_generic no_assms (read_bool arg))
 | 
| 55627 | 449  | 
  | parse_test_param ("expect", [arg]) = apsnd (map_test_params (apsnd (K (read_expectation arg))))
 | 
| 
43881
 
cabe74eab19a
changing parser in quickcheck to activate and deactivate the testers
 
bulwahn 
parents: 
43879 
diff
changeset
 | 
450  | 
  | parse_test_param ("report", [arg]) = apsnd (Config.put_generic report (read_bool arg))
 | 
| 55627 | 451  | 
  | parse_test_param ("genuine_only", [arg]) =
 | 
452  | 
apsnd (Config.put_generic genuine_only (read_bool arg))  | 
|
453  | 
  | parse_test_param ("abort_potential", [arg]) =
 | 
|
454  | 
apsnd (Config.put_generic abort_potential (read_bool arg))  | 
|
| 
43881
 
cabe74eab19a
changing parser in quickcheck to activate and deactivate the testers
 
bulwahn 
parents: 
43879 
diff
changeset
 | 
455  | 
  | parse_test_param ("quiet", [arg]) = apsnd (Config.put_generic quiet (read_bool arg))
 | 
| 45764 | 456  | 
  | parse_test_param ("verbose", [arg]) = apsnd (Config.put_generic verbose (read_bool arg))
 | 
| 46863 | 457  | 
  | parse_test_param ("tag", [arg]) = apsnd (Config.put_generic tag arg)
 | 
| 55627 | 458  | 
  | parse_test_param ("use_subtype", [arg]) =
 | 
459  | 
apsnd (Config.put_generic use_subtype (read_bool arg))  | 
|
460  | 
  | parse_test_param ("timeout", [arg]) =
 | 
|
461  | 
apsnd (Config.put_generic timeout (read_real arg))  | 
|
462  | 
  | parse_test_param ("finite_types", [arg]) =
 | 
|
463  | 
apsnd (Config.put_generic finite_types (read_bool arg))  | 
|
| 
45449
 
eeffd04cd899
adding option allow_function_inversion to quickcheck options
 
bulwahn 
parents: 
45419 
diff
changeset
 | 
464  | 
  | parse_test_param ("allow_function_inversion", [arg]) =
 | 
| 
 
eeffd04cd899
adding option allow_function_inversion to quickcheck options
 
bulwahn 
parents: 
45419 
diff
changeset
 | 
465  | 
apsnd (Config.put_generic allow_function_inversion (read_bool arg))  | 
| 41517 | 466  | 
  | parse_test_param ("finite_type_size", [arg]) =
 | 
| 55627 | 467  | 
apsnd (Config.put_generic finite_type_size (read_nat arg))  | 
| 51302 | 468  | 
| parse_test_param (name, _) =  | 
| 59437 | 469  | 
(fn (testers, context) =>  | 
470  | 
if valid_tester_name context name then  | 
|
471  | 
(insert (op =) name testers, context)  | 
|
| 51302 | 472  | 
        else error ("Unknown tester or test parameter: " ^ name));
 | 
| 28315 | 473  | 
|
| 
43881
 
cabe74eab19a
changing parser in quickcheck to activate and deactivate the testers
 
bulwahn 
parents: 
43879 
diff
changeset
 | 
474  | 
fun parse_test_param_inst (name, arg) ((insts, eval_terms), (testers, ctxt)) =  | 
| 51302 | 475  | 
(case try (Proof_Context.read_typ ctxt) name of  | 
476  | 
SOME (TFree (v, _)) =>  | 
|
477  | 
((AList.update (op =) (v, Proof_Context.read_typ ctxt (the_single arg)) insts, eval_terms),  | 
|
478  | 
(testers, ctxt))  | 
|
479  | 
| NONE =>  | 
|
480  | 
(case name of  | 
|
481  | 
"eval" => ((insts, eval_terms @ map (Syntax.read_term ctxt) arg), (testers, ctxt))  | 
|
482  | 
| _ =>  | 
|
483  | 
((insts, eval_terms),  | 
|
| 55630 | 484  | 
let  | 
485  | 
val (testers', Context.Proof ctxt') =  | 
|
486  | 
parse_test_param (name, arg) (testers, Context.Proof ctxt);  | 
|
487  | 
in (testers', ctxt') end)));  | 
|
| 28309 | 488  | 
|
| 55627 | 489  | 
fun quickcheck_params_cmd args =  | 
490  | 
Context.theory_map  | 
|
| 55629 | 491  | 
(fn context => uncurry set_active_testers (fold parse_test_param args ([], context)));  | 
| 41517 | 492  | 
|
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
493  | 
fun check_expectation state results =  | 
| 51302 | 494  | 
if is_some results andalso expect (Proof.context_of state) = No_Counterexample then  | 
495  | 
error "quickcheck expected to find no counterexample but found one"  | 
|
496  | 
else if is_none results andalso expect (Proof.context_of state) = Counterexample then  | 
|
497  | 
error "quickcheck expected to find a counterexample but did not find one"  | 
|
498  | 
else ();  | 
|
| 
42088
 
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
 
bulwahn 
parents: 
42087 
diff
changeset
 | 
499  | 
|
| 
35378
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35324 
diff
changeset
 | 
500  | 
fun gen_quickcheck args i state =  | 
| 
40644
 
0850a2a16dce
changed old-style quickcheck configurations to new Config.T configurations
 
bulwahn 
parents: 
40643 
diff
changeset
 | 
501  | 
state  | 
| 
43881
 
cabe74eab19a
changing parser in quickcheck to activate and deactivate the testers
 
bulwahn 
parents: 
43879 
diff
changeset
 | 
502  | 
|> Proof.map_context_result (fn ctxt =>  | 
| 
 
cabe74eab19a
changing parser in quickcheck to activate and deactivate the testers
 
bulwahn 
parents: 
43879 
diff
changeset
 | 
503  | 
apsnd (fn (testers, ctxt) => Context.proof_map (set_active_testers testers) ctxt)  | 
| 
 
cabe74eab19a
changing parser in quickcheck to activate and deactivate the testers
 
bulwahn 
parents: 
43879 
diff
changeset
 | 
504  | 
(fold parse_test_param_inst args (([], []), ([], ctxt))))  | 
| 55627 | 505  | 
|> (fn ((insts, eval_terms), state') =>  | 
506  | 
test_goal (true, true) (insts, eval_terms) i state'  | 
|
507  | 
|> tap (check_expectation state')  | 
|
508  | 
|> rpair state');  | 
|
| 32297 | 509  | 
|
| 
43879
 
c8308a8faf9f
enabling parallel execution of testers but removing more informative quickcheck output
 
bulwahn 
parents: 
43878 
diff
changeset
 | 
510  | 
fun quickcheck args i state =  | 
| 51302 | 511  | 
Option.map (the o get_first counterexample_of) (fst (gen_quickcheck args i state));  | 
| 
35378
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35324 
diff
changeset
 | 
512  | 
|
| 
60190
 
906de96ba68a
allow diagnostic proof commands with skip_proofs;
 
wenzelm 
parents: 
60094 
diff
changeset
 | 
513  | 
fun quickcheck_cmd args i st =  | 
| 
 
906de96ba68a
allow diagnostic proof commands with skip_proofs;
 
wenzelm 
parents: 
60094 
diff
changeset
 | 
514  | 
gen_quickcheck args i (Toplevel.proof_of st)  | 
| 46863 | 515  | 
|> apfst (Option.map (the o get_first response_of))  | 
| 51302 | 516  | 
|> (fn (r, state) =>  | 
| 58843 | 517  | 
writeln (Pretty.string_of  | 
| 51302 | 518  | 
(pretty_counterex (Proof.context_of state) false r)));  | 
| 28309 | 519  | 
|
| 41517 | 520  | 
val parse_arg =  | 
521  | 
Parse.name --  | 
|
| 67149 | 522  | 
(Scan.optional (\<^keyword>\<open>=\<close> |--  | 
| 41517 | 523  | 
(((Parse.name || Parse.float_number) >> single) ||  | 
| 67149 | 524  | 
(\<^keyword>\<open>[\<close> |-- Parse.list1 Parse.name --| \<^keyword>\<open>]\<close>))) ["true"]);  | 
| 28309 | 525  | 
|
| 41517 | 526  | 
val parse_args =  | 
| 67149 | 527  | 
\<^keyword>\<open>[\<close> |-- Parse.list1 parse_arg --| \<^keyword>\<open>]\<close> || Scan.succeed [];  | 
| 28336 | 528  | 
|
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36610 
diff
changeset
 | 
529  | 
val _ =  | 
| 67149 | 530  | 
Outer_Syntax.command \<^command_keyword>\<open>quickcheck_params\<close> "set parameters for random testing"  | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36610 
diff
changeset
 | 
531  | 
(parse_args >> (fn args => Toplevel.theory (quickcheck_params_cmd args)));  | 
| 28309 | 532  | 
|
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36610 
diff
changeset
 | 
533  | 
val _ =  | 
| 67149 | 534  | 
Outer_Syntax.command \<^command_keyword>\<open>quickcheck\<close>  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
535  | 
"try to find counterexample for subgoal"  | 
| 60094 | 536  | 
(parse_args -- Scan.optional Parse.nat 1 >>  | 
| 
60190
 
906de96ba68a
allow diagnostic proof commands with skip_proofs;
 
wenzelm 
parents: 
60094 
diff
changeset
 | 
537  | 
(fn (args, i) => Toplevel.keep_proof (quickcheck_cmd args i)));  | 
| 28309 | 538  | 
|
| 51302 | 539  | 
|
| 74508 | 540  | 
(* 'try' setup *)  | 
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
541  | 
|
| 74508 | 542  | 
val _ =  | 
543  | 
Try.tool_setup  | 
|
544  | 
   {name = quickcheckN, weight = 20, auto_option = \<^system_option>\<open>auto_quickcheck\<close>,
 | 
|
545  | 
body = fn auto => fn state =>  | 
|
546  | 
let  | 
|
547  | 
val ctxt = Proof.context_of state;  | 
|
548  | 
val i = 1;  | 
|
549  | 
val res =  | 
|
550  | 
state  | 
|
551  | 
|> Proof.map_context (Config.put report false #> Config.put quiet true)  | 
|
552  | 
|> try (test_goal (false, false) ([], []) i);  | 
|
553  | 
in  | 
|
554  | 
(case res of  | 
|
555  | 
NONE => (unknownN, [])  | 
|
556  | 
| SOME results =>  | 
|
557  | 
let  | 
|
558  | 
val msg =  | 
|
559  | 
Pretty.string_of  | 
|
560  | 
(pretty_counterex ctxt auto (Option.map (the o get_first response_of) results))  | 
|
561  | 
in  | 
|
562  | 
if is_some results then (genuineN, if auto then [msg] else (writeln msg; []))  | 
|
563  | 
else (noneN, [])  | 
|
564  | 
end)  | 
|
565  | 
end  | 
|
566  | 
|> `(fn (outcome_code, _) => outcome_code = genuineN)};  | 
|
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
567  | 
|
| 28315 | 568  | 
end;  |