|
31260
|
1 |
(* Author: Florian Haftmann, TU Muenchen
|
|
|
2 |
|
|
|
3 |
Quickcheck generators for various types.
|
|
|
4 |
*)
|
|
|
5 |
|
|
|
6 |
signature QUICKCHECK_GENERATORS =
|
|
|
7 |
sig
|
|
|
8 |
val compile_generator_expr: theory -> term -> int -> term list option
|
|
|
9 |
type seed = Random_Engine.seed
|
|
|
10 |
val random_fun: typ -> typ -> ('a -> 'a -> bool) -> ('a -> term)
|
|
|
11 |
-> (seed -> ('b * (unit -> term)) * seed) -> (seed -> seed * seed)
|
|
|
12 |
-> seed -> (('a -> 'b) * (unit -> Term.term)) * seed
|
|
|
13 |
val ensure_random_typecopy: string -> theory -> theory
|
|
|
14 |
val eval_ref: (unit -> int -> int * int -> term list option * (int * int)) option ref
|
|
|
15 |
val setup: theory -> theory
|
|
|
16 |
end;
|
|
|
17 |
|
|
|
18 |
structure Quickcheck_Generators : QUICKCHECK_GENERATORS =
|
|
|
19 |
struct
|
|
|
20 |
|
|
|
21 |
(** building and compiling generator expressions **)
|
|
|
22 |
|
|
|
23 |
val eval_ref : (unit -> int -> int * int -> term list option * (int * int)) option ref = ref NONE;
|
|
|
24 |
|
|
|
25 |
val target = "Quickcheck";
|
|
|
26 |
|
|
|
27 |
fun mk_generator_expr thy prop tys =
|
|
|
28 |
let
|
|
|
29 |
val bound_max = length tys - 1;
|
|
|
30 |
val bounds = map_index (fn (i, ty) =>
|
|
|
31 |
(2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) tys;
|
|
|
32 |
val result = list_comb (prop, map (fn (i, _, _, _) => Bound i) bounds);
|
|
|
33 |
val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) bounds);
|
|
|
34 |
val check = @{term "If :: bool => term list option => term list option => term list option"}
|
|
|
35 |
$ result $ @{term "None :: term list option"} $ (@{term "Some :: term list => term list option "} $ terms);
|
|
|
36 |
val return = @{term "Pair :: term list option => Random.seed => term list option * Random.seed"};
|
|
|
37 |
fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT);
|
|
|
38 |
fun mk_termtyp ty = HOLogic.mk_prodT (ty, @{typ "unit => term"});
|
|
|
39 |
fun mk_scomp T1 T2 sT f g = Const (@{const_name scomp},
|
|
|
40 |
liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g;
|
|
|
41 |
fun mk_split ty = Sign.mk_const thy
|
|
|
42 |
(@{const_name split}, [ty, @{typ "unit => term"}, liftT @{typ "term list option"} @{typ Random.seed}]);
|
|
|
43 |
fun mk_scomp_split ty t t' =
|
|
|
44 |
mk_scomp (mk_termtyp ty) @{typ "term list option"} @{typ Random.seed} t
|
|
|
45 |
(mk_split ty $ Abs ("", ty, Abs ("", @{typ "unit => term"}, t')));
|
|
|
46 |
fun mk_bindclause (_, _, i, ty) = mk_scomp_split ty
|
|
|
47 |
(Sign.mk_const thy (@{const_name random}, [ty]) $ Bound i);
|
|
|
48 |
in Abs ("n", @{typ code_numeral}, fold_rev mk_bindclause bounds (return $ check)) end;
|
|
|
49 |
|
|
|
50 |
fun compile_generator_expr thy t =
|
|
|
51 |
let
|
|
|
52 |
val tys = (map snd o fst o strip_abs) t;
|
|
|
53 |
val t' = mk_generator_expr thy t tys;
|
|
|
54 |
val f = Code_ML.eval (SOME target) ("Quickcheck_Generators.eval_ref", eval_ref)
|
|
|
55 |
(fn proc => fn g => fn s => g s #>> (Option.map o map) proc) thy t' [];
|
|
|
56 |
in f #> Random_Engine.run end;
|
|
|
57 |
|
|
|
58 |
|
|
|
59 |
(** typ "'a => 'b" **)
|
|
|
60 |
|
|
|
61 |
type seed = Random_Engine.seed;
|
|
|
62 |
|
|
|
63 |
fun random_fun (T1 : typ) (T2 : typ) (eq : 'a -> 'a -> bool) (term_of : 'a -> term)
|
|
|
64 |
(random : seed -> ('b * (unit -> term)) * seed)
|
|
|
65 |
(random_split : seed -> seed * seed)
|
|
|
66 |
(seed : seed) =
|
|
|
67 |
let
|
|
|
68 |
val (seed', seed'') = random_split seed;
|
|
|
69 |
val state = ref (seed', [], Const (@{const_name undefined}, T1 --> T2));
|
|
|
70 |
val fun_upd = Const (@{const_name fun_upd},
|
|
|
71 |
(T1 --> T2) --> T1 --> T2 --> T1 --> T2);
|
|
|
72 |
fun random_fun' x =
|
|
|
73 |
let
|
|
|
74 |
val (seed, fun_map, f_t) = ! state;
|
|
|
75 |
in case AList.lookup (uncurry eq) fun_map x
|
|
|
76 |
of SOME y => y
|
|
|
77 |
| NONE => let
|
|
|
78 |
val t1 = term_of x;
|
|
|
79 |
val ((y, t2), seed') = random seed;
|
|
|
80 |
val fun_map' = (x, y) :: fun_map;
|
|
|
81 |
val f_t' = fun_upd $ f_t $ t1 $ t2 ();
|
|
|
82 |
val _ = state := (seed', fun_map', f_t');
|
|
|
83 |
in y end
|
|
|
84 |
end;
|
|
|
85 |
fun term_fun' () = #3 (! state);
|
|
|
86 |
in ((random_fun', term_fun'), seed'') end;
|
|
|
87 |
|
|
|
88 |
|
|
|
89 |
(** type copies **)
|
|
|
90 |
|
|
|
91 |
fun mk_random_typecopy tyco vs constr typ thy =
|
|
|
92 |
let
|
|
|
93 |
val Ts = map TFree vs;
|
|
|
94 |
val T = Type (tyco, Ts);
|
|
|
95 |
fun mk_termifyT T = HOLogic.mk_prodT (T, @{typ "unit => term"})
|
|
|
96 |
val Ttm = mk_termifyT T;
|
|
|
97 |
val typtm = mk_termifyT typ;
|
|
|
98 |
fun mk_const c Ts = Const (c, Sign.const_instance thy (c, Ts));
|
|
|
99 |
fun mk_random T = mk_const @{const_name random} [T];
|
|
|
100 |
val size = @{term "j::code_numeral"};
|
|
|
101 |
val v = "x";
|
|
|
102 |
val t_v = Free (v, typtm);
|
|
|
103 |
val t_constr = mk_const constr Ts;
|
|
|
104 |
val lhs = mk_random T $ size;
|
|
|
105 |
val rhs = HOLogic.mk_ST [(((mk_random typ) $ size, @{typ Random.seed}), SOME (v, typtm))]
|
|
|
106 |
(HOLogic.mk_return Ttm @{typ Random.seed}
|
|
|
107 |
(mk_const "Code_Eval.valapp" [typ, T]
|
|
|
108 |
$ HOLogic.mk_prod (t_constr, Abs ("u", @{typ unit}, HOLogic.reflect_term t_constr)) $ t_v))
|
|
|
109 |
@{typ Random.seed} (SOME Ttm, @{typ Random.seed});
|
|
|
110 |
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
|
|
|
111 |
in
|
|
|
112 |
thy
|
|
|
113 |
|> TheoryTarget.instantiation ([tyco], vs, @{sort random})
|
|
|
114 |
|> `(fn lthy => Syntax.check_term lthy eq)
|
|
|
115 |
|-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq)))
|
|
|
116 |
|> snd
|
|
|
117 |
|> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
|
|
|
118 |
end;
|
|
|
119 |
|
|
|
120 |
fun ensure_random_typecopy tyco thy =
|
|
|
121 |
let
|
|
|
122 |
val SOME { vs = raw_vs, constr, typ = raw_typ, ... } =
|
|
|
123 |
TypecopyPackage.get_info thy tyco;
|
|
|
124 |
val constrain = curry (Sorts.inter_sort (Sign.classes_of thy));
|
|
|
125 |
val typ = map_atyps (fn TFree (v, sort) =>
|
|
|
126 |
TFree (v, constrain sort @{sort random})) raw_typ;
|
|
|
127 |
val vs' = Term.add_tfreesT typ [];
|
|
|
128 |
val vs = map (fn (v, sort) =>
|
|
|
129 |
(v, the_default (constrain sort @{sort typerep}) (AList.lookup (op =) vs' v))) raw_vs;
|
|
|
130 |
val do_inst = Sign.of_sort thy (typ, @{sort random});
|
|
|
131 |
in if do_inst then mk_random_typecopy tyco vs constr typ thy else thy end;
|
|
|
132 |
|
|
|
133 |
|
|
|
134 |
(** datatypes **)
|
|
|
135 |
|
|
|
136 |
(* still under construction *)
|
|
|
137 |
|
|
|
138 |
|
|
|
139 |
(** setup **)
|
|
|
140 |
|
|
|
141 |
val setup = Code_Target.extend_target (target, (Code_ML.target_Eval, K I))
|
|
|
142 |
#> Quickcheck.add_generator ("code", compile_generator_expr o ProofContext.theory_of)
|
|
|
143 |
#> TypecopyPackage.interpretation ensure_random_typecopy;
|
|
|
144 |
|
|
|
145 |
end;
|