author | haftmann |
Fri, 25 Sep 2009 09:50:31 +0200 | |
changeset 32705 | 04ce6bb14d85 |
parent 32672 | 90f3ce5d27ae |
child 32740 | 9dd0a2f83429 |
permissions | -rw-r--r-- |
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
1 |
(* Author: Lukas Bulwahn, TU Muenchen |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
2 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
3 |
A quickcheck generator based on the predicate compiler |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
4 |
*) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
5 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
6 |
signature PRED_COMPILE_QUICKCHECK = |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
7 |
sig |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
8 |
val quickcheck : Proof.context -> term -> int -> term list option |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
9 |
val test_ref : ((unit -> int -> int * int -> term list Predicate.pred * (int * int)) option) ref |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
10 |
end; |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
11 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
12 |
structure Pred_Compile_Quickcheck : PRED_COMPILE_QUICKCHECK = |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
13 |
struct |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
14 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
15 |
val test_ref = ref (NONE : (unit -> int -> int * int -> term list Predicate.pred * (int * int)) option) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
16 |
val target = "Quickcheck" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
17 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
18 |
fun dest_compfuns (Predicate_Compile_Core.CompilationFuns funs) = funs |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
19 |
val mk_predT = #mk_predT (dest_compfuns Predicate_Compile_Core.pred_compfuns) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
20 |
val mk_rpredT = #mk_predT (dest_compfuns Predicate_Compile_Core.rpred_compfuns) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
21 |
val mk_return = #mk_single (dest_compfuns Predicate_Compile_Core.rpred_compfuns) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
22 |
val mk_bind = #mk_bind (dest_compfuns Predicate_Compile_Core.rpred_compfuns) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
23 |
val lift_pred = #lift_pred (dest_compfuns Predicate_Compile_Core.rpred_compfuns) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
24 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
25 |
fun mk_split_lambda [] t = Abs ("u", HOLogic.unitT, t) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
26 |
| mk_split_lambda [x] t = lambda x t |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
27 |
| mk_split_lambda xs t = |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
28 |
let |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
29 |
fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t)) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
30 |
| mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t)) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
31 |
in |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
32 |
mk_split_lambda' xs t |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
33 |
end; |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
34 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
35 |
fun strip_imp_prems (Const("op -->", _) $ A $ B) = A :: strip_imp_prems B |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
36 |
| strip_imp_prems _ = []; |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
37 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
38 |
fun strip_imp_concl (Const("op -->", _) $ A $ B) = strip_imp_concl B |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
39 |
| strip_imp_concl A = A : term; |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
40 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
41 |
fun strip_horn A = (strip_imp_prems A, strip_imp_concl A); |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
42 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
43 |
fun quickcheck ctxt t = |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
44 |
let |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
45 |
val _ = tracing ("Starting quickcheck with " ^ (Syntax.string_of_term ctxt t)) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
46 |
val ctxt' = ProofContext.theory (Context.copy_thy) ctxt |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
47 |
val thy = (ProofContext.theory_of ctxt') |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
48 |
val (vs, t') = strip_abs t |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
49 |
val vs' = Variable.variant_frees ctxt' [] vs |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
50 |
val t'' = subst_bounds (map Free (rev vs'), t') |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
51 |
val (prems, concl) = strip_horn t'' |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
52 |
val constname = "pred_compile_quickcheck" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
53 |
val full_constname = Sign.full_bname thy constname |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
54 |
val constT = map snd vs' ---> @{typ bool} |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
55 |
val thy' = Sign.add_consts_i [(Binding.name constname, constT, NoSyn)] thy |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
56 |
val t = Logic.list_implies |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
57 |
(map HOLogic.mk_Trueprop (prems @ [HOLogic.mk_not concl]), |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
58 |
HOLogic.mk_Trueprop (list_comb (Const (full_constname, constT), map Free vs'))) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
59 |
val tac = fn {...} => setmp quick_and_dirty true (SkipProof.cheat_tac thy') |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
60 |
val intro = Goal.prove (ProofContext.init thy') (map fst vs') [] t tac |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
61 |
val _ = tracing (Display.string_of_thm ctxt' intro) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
62 |
val thy'' = thy' |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
63 |
|> Context.theory_map (Predicate_Compile_Preproc_Const_Defs.add_thm intro) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
64 |
|> Predicate_Compile.preprocess full_constname |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
65 |
|> Predicate_Compile_Core.add_equations [full_constname] |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
66 |
|> Predicate_Compile_Core.add_sizelim_equations [full_constname] |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
67 |
|> Predicate_Compile_Core.add_quickcheck_equations [full_constname] |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
68 |
val sizelim_modes = Predicate_Compile_Core.sizelim_modes_of thy'' full_constname |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
69 |
val modes = Predicate_Compile_Core.generator_modes_of thy'' full_constname |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
70 |
val prog = |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
71 |
if member (op =) modes ([], []) then |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
72 |
let |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
73 |
val name = Predicate_Compile_Core.generator_name_of thy'' full_constname ([], []) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
74 |
val T = @{typ code_numeral} --> (mk_rpredT (HOLogic.mk_tupleT (map snd vs'))) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
75 |
in Const (name, T) $ Bound 0 end |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
76 |
else if member (op =) sizelim_modes ([], []) then |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
77 |
let |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
78 |
val name = Predicate_Compile_Core.sizelim_function_name_of thy'' full_constname ([], []) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
79 |
val T = @{typ code_numeral} --> (mk_predT (HOLogic.mk_tupleT (map snd vs'))) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
80 |
in lift_pred (Const (name, T) $ Bound 0) end |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
81 |
else error "Predicate Compile Quickcheck failed" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
82 |
val qc_term = Abs ("size", @{typ code_numeral}, mk_bind (prog, |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
83 |
mk_split_lambda (map Free vs') (mk_return (HOLogic.mk_list @{typ term} |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
84 |
(map2 HOLogic.mk_term_of (map snd vs') (map Free vs')))))) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
85 |
val _ = tracing (Syntax.string_of_term ctxt' qc_term) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
86 |
val compile = Code_ML.eval (SOME target) ("Pred_Compile_Quickcheck.test_ref", test_ref) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
87 |
(fn proc => fn g => fn s => g s #>> (Predicate.map o map) proc) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
88 |
thy'' qc_term [] |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
89 |
in |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
90 |
((compile #> Random_Engine.run) #> (Option.map fst o Predicate.yield)) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
91 |
end |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
92 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
diff
changeset
|
93 |
end; |