| author | blanchet | 
| Thu, 22 May 2014 03:29:35 +0200 | |
| changeset 57053 | 46000c075d07 | 
| parent 51143 | 0a2371e7ced3 | 
| child 58812 | 5a9a2d3b9f8b | 
| permissions | -rw-r--r-- | 
| 35953 | 1  | 
(* Author: Lukas Bulwahn, TU Muenchen *)  | 
2  | 
||
3  | 
header {* A Prototype of Quickcheck based on the Predicate Compiler *}
 | 
|
4  | 
||
5  | 
theory Predicate_Compile_Quickcheck  | 
|
6  | 
imports Main Predicate_Compile_Alternative_Defs  | 
|
7  | 
begin  | 
|
8  | 
||
| 48891 | 9  | 
ML_file "../Tools/Predicate_Compile/predicate_compile_quickcheck.ML"  | 
10  | 
||
| 
43886
 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 
bulwahn 
parents: 
39252 
diff
changeset
 | 
11  | 
setup {* Predicate_Compile_Quickcheck.setup *}
 | 
| 
 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 
bulwahn 
parents: 
39252 
diff
changeset
 | 
12  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48891 
diff
changeset
 | 
13  | 
end  |