| author | wenzelm | 
| Thu, 22 Sep 2016 11:25:27 +0200 | |
| changeset 63936 | b87784e19a77 | 
| parent 63167 | 0909deb8059b | 
| child 69597 | ff784d5a5bfb | 
| permissions | -rw-r--r-- | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
1  | 
theory Quickcheck_Interfaces  | 
| 
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
2  | 
imports Main  | 
| 
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
3  | 
begin  | 
| 
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
4  | 
|
| 63167 | 5  | 
subsection \<open>Checking a single proposition (TODO)\<close>  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
6  | 
|
| 63167 | 7  | 
subsection \<open>Checking multiple propositions in one batch\<close>  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
8  | 
|
| 63167 | 9  | 
text \<open>  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
10  | 
|
| 
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
11  | 
First, this requires to setup special generators for all datatypes via the following command.  | 
| 63167 | 12  | 
\<close>  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
13  | 
|
| 58813 | 14  | 
setup Exhaustive_Generators.setup_bounded_forall_datatype_interpretation  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
15  | 
|
| 63167 | 16  | 
text \<open>  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
17  | 
Now, the function Quickcheck.mk_batch_validator : Proof.context -> term list -> (int -> bool) list option  | 
| 
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
18  | 
takes formulas of type bool with free variables, and returns a list of testing functions.  | 
| 63167 | 19  | 
\<close>  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
20  | 
|
| 63167 | 21  | 
ML \<open>  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
22  | 
val SOME testers = Quickcheck.mk_batch_validator @{context}
 | 
| 
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
23  | 
  [@{term "x = (1 :: nat)"}, @{term "x = (0 :: nat)"}, @{term "x <= (5 :: nat)"}, @{term "0 \<le> (x :: nat)"}]
 | 
| 63167 | 24  | 
\<close>  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
25  | 
|
| 63167 | 26  | 
text \<open>  | 
| 58148 | 27  | 
It is up to the user with which strategy the conjectures should be tested.  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
28  | 
For example, one could check all conjectures up to a given size, and check the different conjectures in sequence.  | 
| 
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
29  | 
This is implemented by:  | 
| 63167 | 30  | 
\<close>  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
31  | 
|
| 63167 | 32  | 
ML \<open>  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
33  | 
fun check_upto f i j = if i > j then true else f i andalso check_upto f (i + 1) j  | 
| 63167 | 34  | 
\<close>  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
35  | 
|
| 63167 | 36  | 
ML_val \<open>  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
37  | 
map (fn test => check_upto test 0 1) testers  | 
| 63167 | 38  | 
\<close>  | 
39  | 
ML_val \<open>  | 
|
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
40  | 
map (fn test => check_upto test 0 2) testers  | 
| 63167 | 41  | 
\<close>  | 
42  | 
ML_val \<open>  | 
|
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
43  | 
map (fn test => check_upto test 0 3) testers  | 
| 63167 | 44  | 
\<close>  | 
45  | 
ML_val \<open>  | 
|
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
46  | 
map (fn test => check_upto test 0 7) testers  | 
| 63167 | 47  | 
\<close>  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
48  | 
|
| 63167 | 49  | 
text \<open>Note that all conjectures must be executable to obtain the testers with the function above.\<close>  | 
| 
48179
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
50  | 
|
| 
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
18461f745b4a
adding some minimal documentation and an example of quickcheck's interfaces
 
bulwahn 
parents:  
diff
changeset
 | 
52  | 
end  |