| author | Christian Sternagel | 
| Thu, 30 Aug 2012 13:06:04 +0900 | |
| changeset 49088 | 5cd8b4426a57 | 
| parent 48891 | c0eafbd55de3 | 
| child 51143 | 0a2371e7ced3 | 
| 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: 
39252diff
changeset | 11 | setup {* Predicate_Compile_Quickcheck.setup *}
 | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
39252diff
changeset | 12 | |
| 35953 | 13 | end |