src/HOL/Library/Predicate_Compile_Quickcheck.thy
author bulwahn
Thu, 09 Sep 2010 16:43:57 +0200
changeset 39252 8f176e575a49
parent 36026 276ebec72082
child 43886 bf068e758783
permissions -rw-r--r--
changing the container for the quickcheck options to a generic data
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     1
(* Author: Lukas Bulwahn, TU Muenchen *)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     2
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     3
header {* A Prototype of Quickcheck based on the Predicate Compiler *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     4
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     5
theory Predicate_Compile_Quickcheck
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     6
imports Main Predicate_Compile_Alternative_Defs
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     7
uses "../Tools/Predicate_Compile/predicate_compile_quickcheck.ML"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     8
begin
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     9
39252
8f176e575a49 changing the container for the quickcheck options to a generic data
bulwahn
parents: 36026
diff changeset
    10
setup {* Context.theory_map (Quickcheck.add_generator ("predicate_compile_wo_ff", Predicate_Compile_Quickcheck.quickcheck_compile_term
8f176e575a49 changing the container for the quickcheck options to a generic data
bulwahn
parents: 36026
diff changeset
    11
  Predicate_Compile_Aux.New_Pos_Random_DSeq false true 4)) *}
8f176e575a49 changing the container for the quickcheck options to a generic data
bulwahn
parents: 36026
diff changeset
    12
setup {* Context.theory_map (Quickcheck.add_generator ("predicate_compile_ff_fs",
8f176e575a49 changing the container for the quickcheck options to a generic data
bulwahn
parents: 36026
diff changeset
    13
  Predicate_Compile_Quickcheck.quickcheck_compile_term Predicate_Compile_Aux.New_Pos_Random_DSeq true true 4)) *}
8f176e575a49 changing the container for the quickcheck options to a generic data
bulwahn
parents: 36026
diff changeset
    14
setup {* Context.theory_map (Quickcheck.add_generator ("predicate_compile_ff_nofs",
8f176e575a49 changing the container for the quickcheck options to a generic data
bulwahn
parents: 36026
diff changeset
    15
  Predicate_Compile_Quickcheck.quickcheck_compile_term Predicate_Compile_Aux.New_Pos_Random_DSeq true false 4)) *}
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    16
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    17
end