src/HOL/Quickcheck_Examples/Quickcheck_Interfaces.thy
author wenzelm
Tue, 25 Sep 2012 20:28:47 +0200
changeset 49565 ea4308b7ef0f
parent 48179 18461f745b4a
child 51272 9c8d63b4b6be
permissions -rw-r--r--
ML support for generic graph display, with browser and graphview backends (via print modes); reverse graph layout (again), according to the graph orientation provided by ML; simplified scala types -- eliminated unused type parameters; more explicit Model.Info, Model.Graph; renamed isabelle.graphview.Graphview_Frame to isabelle.graphview.Frame in accordance to file name; removed obsolete Graph_XML and Tooltips; tuned graphview command line; more generous JVM resources via GRAPHVIEW_JAVA_OPTIONS;

theory Quickcheck_Interfaces
imports Main
begin

subsection {* Checking a single proposition (TODO) *}

subsection {* Checking multiple propositions in one batch *}

text {*

First, this requires to setup special generators for all datatypes via the following command.
*}

setup {* Exhaustive_Generators.setup_bounded_forall_datatype_interpretation *}

text {*
Now, the function Quickcheck.mk_batch_validator : Proof.context -> term list -> (int -> bool) list option
takes formulas of type bool with free variables, and returns a list of testing functions.
*}

ML {*
val SOME testers = Quickcheck.mk_batch_validator @{context}
  [@{term "x = (1 :: nat)"}, @{term "x = (0 :: nat)"}, @{term "x <= (5 :: nat)"}, @{term "0 \<le> (x :: nat)"}]
*}

text {*
It is upto the user with which strategy the conjectures should be tested.
For example, one could check all conjectures up to a given size, and check the different conjectures in sequence.
This is implemented by:
*}

ML {*
fun check_upto f i j = if i > j then true else f i andalso check_upto f (i + 1) j
*}

ML {* 
  map (fn test => check_upto test 0 1) testers
*}
ML {* 
  map (fn test => check_upto test 0 2) testers
*}
ML {* 
  map (fn test => check_upto test 0 3) testers
*}
ML {* 
  map (fn test => check_upto test 0 7) testers
*}

text {* Note that all conjectures must be executable to obtain the testers with the function above. *}


end