| author | wenzelm | 
| Sun, 27 Mar 2011 15:01:47 +0200 | |
| changeset 42130 | e10e2cce85c8 | 
| parent 42029 | da92153d6dff | 
| child 42159 | 234ec7011e5d | 
| permissions | -rw-r--r-- | 
| 33084 | 1 | |
| 2 | header {* Install quickcheck of SML code generator *}
 | |
| 3 | ||
| 4 | theory SML_Quickcheck | |
| 5 | imports Main | |
| 6 | begin | |
| 7 | ||
| 8 | setup {*
 | |
| 37390 
8781d80026fc
moved inductive_codegen to place where product type is available; tuned structure name
 haftmann parents: 
33771diff
changeset | 9 | Inductive_Codegen.quickcheck_setup #> | 
| 40919 | 10 |   Context.theory_map (Quickcheck.add_generator ("SML",
 | 
| 42029 | 11 | fn ctxt => fn (t, eval_terms) => | 
| 40919 | 12 | let | 
| 13 | val test_fun = Codegen.test_term ctxt t | |
| 14 | val iterations = Config.get ctxt Quickcheck.iterations | |
| 15 | fun iterate size 0 = NONE | |
| 16 | | iterate size j = | |
| 17 | let | |
| 18 | val result = test_fun size handle Match => | |
| 19 | (if Config.get ctxt Quickcheck.quiet then () | |
| 20 | else warning "Exception Match raised during quickcheck"; NONE) | |
| 21 | in | |
| 22 | case result of NONE => iterate size (j - 1) | SOME q => SOME q | |
| 23 | end | |
| 24 | in fn size => (iterate size iterations, NONE) end)) | |
| 33084 | 25 | *} | 
| 26 | ||
| 27 | end |