src/HOL/Quickcheck_Examples/Quickcheck_Interfaces.thy
author wenzelm
Sat, 13 Aug 2022 18:06:30 +0200
changeset 75848 9e4c0aaa30aa
parent 69597 ff784d5a5bfb
permissions -rw-r--r--
clarified signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
     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
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
     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
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
     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
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    12
\<close>
48179
18461f745b4a adding some minimal documentation and an example of quickcheck's interfaces
bulwahn
parents:
diff changeset
    13
58813
wenzelm
parents: 58148
diff changeset
    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
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    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
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    19
\<close>
48179
18461f745b4a adding some minimal documentation and an example of quickcheck's interfaces
bulwahn
parents:
diff changeset
    20
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    21
ML \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63167
diff changeset
    22
val SOME testers = Quickcheck.mk_batch_validator \<^context>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63167
diff changeset
    23
  [\<^term>\<open>x = (1 :: nat)\<close>, \<^term>\<open>x = (0 :: nat)\<close>, \<^term>\<open>x <= (5 :: nat)\<close>, \<^term>\<open>0 \<le> (x :: nat)\<close>]
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    24
\<close>
48179
18461f745b4a adding some minimal documentation and an example of quickcheck's interfaces
bulwahn
parents:
diff changeset
    25
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    26
text \<open>
58148
9764b994a421 use 'datatype_new' in Quickcheck examples
blanchet
parents: 51272
diff changeset
    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
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    30
\<close>
48179
18461f745b4a adding some minimal documentation and an example of quickcheck's interfaces
bulwahn
parents:
diff changeset
    31
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    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
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    34
\<close>
48179
18461f745b4a adding some minimal documentation and an example of quickcheck's interfaces
bulwahn
parents:
diff changeset
    35
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    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
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    38
\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    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
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    41
\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    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
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    44
\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    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
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    47
\<close>
48179
18461f745b4a adding some minimal documentation and an example of quickcheck's interfaces
bulwahn
parents:
diff changeset
    48
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58813
diff changeset
    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