author | bulwahn |
Thu, 01 Dec 2011 22:14:35 +0100 | |
changeset 45719 | 39c52a820f6e |
parent 45718 | 8979b2463fc8 |
child 45721 | d1fb55c2ed65 |
permissions | -rw-r--r-- |
41923
f05fc0711bc7
renaming signatures and structures; correcting header
bulwahn
parents:
41921
diff
changeset
|
1 |
(* Title: HOL/Tools/Quickcheck/random_generators.ML |
37744 | 2 |
Author: Florian Haftmann, TU Muenchen |
31260 | 3 |
|
41923
f05fc0711bc7
renaming signatures and structures; correcting header
bulwahn
parents:
41921
diff
changeset
|
4 |
Random generators for various types. |
31260 | 5 |
*) |
6 |
||
41923
f05fc0711bc7
renaming signatures and structures; correcting header
bulwahn
parents:
41921
diff
changeset
|
7 |
signature RANDOM_GENERATORS = |
31260 | 8 |
sig |
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) |
|
33243 | 12 |
-> seed -> (('a -> 'b) * (unit -> term)) * seed |
35378
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
13 |
val compile_generator_expr: |
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
14 |
Proof.context -> (term * term list) list -> int list -> term list option * Quickcheck.report option |
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
15 |
val put_counterexample: (unit -> int -> int -> seed -> term list option * seed) |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
16 |
-> Proof.context -> Proof.context |
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
17 |
val put_counterexample_report: (unit -> int -> int -> seed -> (term list option * (bool list * bool)) * seed) |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
18 |
-> Proof.context -> Proof.context |
31260 | 19 |
val setup: theory -> theory |
20 |
end; |
|
21 |
||
41923
f05fc0711bc7
renaming signatures and structures; correcting header
bulwahn
parents:
41921
diff
changeset
|
22 |
structure Random_Generators : RANDOM_GENERATORS = |
31260 | 23 |
struct |
24 |
||
31950 | 25 |
(** abstract syntax **) |
31260 | 26 |
|
31950 | 27 |
fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit => term"}) |
28 |
val size = @{term "i::code_numeral"}; |
|
31984 | 29 |
val size_pred = @{term "(i::code_numeral) - 1"}; |
31950 | 30 |
val size' = @{term "j::code_numeral"}; |
31 |
val seed = @{term "s::Random.seed"}; |
|
31260 | 32 |
|
33 |
||
34 |
(** typ "'a => 'b" **) |
|
35 |
||
36 |
type seed = Random_Engine.seed; |
|
37 |
||
31603
fa30cd74d7d6
revised interpretation combinator for datatype constructions
haftmann
parents:
31595
diff
changeset
|
38 |
fun random_fun T1 T2 eq term_of random random_split seed = |
31260 | 39 |
let |
40 |
val fun_upd = Const (@{const_name fun_upd}, |
|
41 |
(T1 --> T2) --> T1 --> T2 --> T1 --> T2); |
|
32344 | 42 |
val ((y, t2), seed') = random seed; |
43 |
val (seed'', seed''') = random_split seed'; |
|
31933
cd6511035315
proper closures -- imperative programming considered harmful...
haftmann
parents:
31902
diff
changeset
|
44 |
|
32740 | 45 |
val state = Unsynchronized.ref (seed'', [], fn () => Abs ("x", T1, t2 ())); |
31260 | 46 |
fun random_fun' x = |
47 |
let |
|
48 |
val (seed, fun_map, f_t) = ! state; |
|
49 |
in case AList.lookup (uncurry eq) fun_map x |
|
50 |
of SOME y => y |
|
51 |
| NONE => let |
|
52 |
val t1 = term_of x; |
|
53 |
val ((y, t2), seed') = random seed; |
|
54 |
val fun_map' = (x, y) :: fun_map; |
|
31933
cd6511035315
proper closures -- imperative programming considered harmful...
haftmann
parents:
31902
diff
changeset
|
55 |
val f_t' = fn () => fun_upd $ f_t () $ t1 $ t2 (); |
31260 | 56 |
val _ = state := (seed', fun_map', f_t'); |
57 |
in y end |
|
58 |
end; |
|
31933
cd6511035315
proper closures -- imperative programming considered harmful...
haftmann
parents:
31902
diff
changeset
|
59 |
fun term_fun' () = #3 (! state) (); |
32344 | 60 |
in ((random_fun', term_fun'), seed''') end; |
31260 | 61 |
|
62 |
||
63 |
(** datatypes **) |
|
64 |
||
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
65 |
(* definitional scheme for random instances on datatypes *) |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
66 |
|
31611
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
67 |
local |
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
68 |
|
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
69 |
fun dest_ctyp_nth k cT = nth (Thm.dest_ctyp cT) k; |
31611
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
70 |
val eq = Thm.cprop_of @{thm random_aux_rec} |> Thm.dest_arg |> Thm.dest_arg |> Thm.dest_arg; |
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
71 |
val lhs = eq |> Thm.dest_arg1; |
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
72 |
val pt_random_aux = lhs |> Thm.dest_fun; |
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
73 |
val ct_k = lhs |> Thm.dest_arg; |
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
74 |
val pt_rhs = eq |> Thm.dest_arg |> Thm.dest_fun; |
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
75 |
val aT = pt_random_aux |> Thm.ctyp_of_term |> dest_ctyp_nth 1; |
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
76 |
|
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
77 |
val rew_thms = map mk_meta_eq [@{thm code_numeral_zero_minus_one}, |
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
78 |
@{thm Suc_code_numeral_minus_one}, @{thm select_weight_cons_zero}, @{thm beyond_zero}]; |
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
79 |
val rew_ts = map (Logic.dest_equals o Thm.prop_of) rew_thms; |
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
80 |
val rew_ss = HOL_ss addsimps rew_thms; |
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
81 |
|
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
82 |
in |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
83 |
|
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
84 |
fun random_aux_primrec eq lthy = |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
85 |
let |
42361 | 86 |
val thy = Proof_Context.theory_of lthy; |
31611
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
87 |
val ((t_random_aux as Free (random_aux, T)) $ (t_k as Free (v, _)), proto_t_rhs) = |
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
88 |
(HOLogic.dest_eq o HOLogic.dest_Trueprop) eq; |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
89 |
val Type (_, [_, iT]) = T; |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
90 |
val icT = Thm.ctyp_of thy iT; |
31611
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
91 |
val cert = Thm.cterm_of thy; |
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
92 |
val inst = Thm.instantiate_cterm ([(aT, icT)], []); |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
93 |
fun subst_v t' = map_aterms (fn t as Free (w, _) => if v = w then t' else t | t => t); |
31611
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
94 |
val t_rhs = lambda t_k proto_t_rhs; |
31785 | 95 |
val eqs0 = [subst_v @{term "0::code_numeral"} eq, |
96 |
subst_v (@{term "Suc_code_numeral"} $ t_k) eq]; |
|
31611
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
97 |
val eqs1 = map (Pattern.rewrite_term thy rew_ts []) eqs0; |
35166 | 98 |
val ((_, (_, eqs2)), lthy') = Primrec.add_primrec_simple |
33205 | 99 |
[((Binding.conceal (Binding.name random_aux), T), NoSyn)] eqs1 lthy; |
31611
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
100 |
val cT_random_aux = inst pt_random_aux; |
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
101 |
val cT_rhs = inst pt_rhs; |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
102 |
val rule = @{thm random_aux_rec} |
43333
2bdec7f430d3
renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
wenzelm
parents:
43329
diff
changeset
|
103 |
|> Drule.instantiate_normalize ([(aT, icT)], |
31785 | 104 |
[(cT_random_aux, cert t_random_aux), (cT_rhs, cert t_rhs)]); |
105 |
val tac = ALLGOALS (rtac rule) |
|
106 |
THEN ALLGOALS (simp_tac rew_ss) |
|
42361 | 107 |
THEN (ALLGOALS (Proof_Context.fact_tac eqs2)) |
32970
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents:
32740
diff
changeset
|
108 |
val simp = Skip_Proof.prove lthy' [v] [] eq (K tac); |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
109 |
in (simp, lthy') end; |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
110 |
|
31611
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
111 |
end; |
a577f77af93f
explicit instantiation yields considerable speedup
haftmann
parents:
31609
diff
changeset
|
112 |
|
31868 | 113 |
fun random_aux_primrec_multi auxname [eq] lthy = |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
114 |
lthy |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
115 |
|> random_aux_primrec eq |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
116 |
|>> (fn simp => [simp]) |
31868 | 117 |
| random_aux_primrec_multi auxname (eqs as _ :: _ :: _) lthy = |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
118 |
let |
42361 | 119 |
val thy = Proof_Context.theory_of lthy; |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
120 |
val (lhss, rhss) = map_split (HOLogic.dest_eq o HOLogic.dest_Trueprop) eqs; |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
121 |
val (vs, (arg as Free (v, _)) :: _) = map_split (fn (t1 $ t2) => (t1, t2)) lhss; |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
122 |
val Ts = map fastype_of lhss; |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
123 |
val tupleT = foldr1 HOLogic.mk_prodT Ts; |
31868 | 124 |
val aux_lhs = Free ("mutual_" ^ auxname, fastype_of arg --> tupleT) $ arg; |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
125 |
val aux_eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq) |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
126 |
(aux_lhs, foldr1 HOLogic.mk_prod rhss); |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
127 |
fun mk_proj t [T] = [t] |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
128 |
| mk_proj t (Ts as T :: (Ts' as _ :: _)) = |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
129 |
Const (@{const_name fst}, foldr1 HOLogic.mk_prodT Ts --> T) $ t |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
130 |
:: mk_proj (Const (@{const_name snd}, |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
131 |
foldr1 HOLogic.mk_prodT Ts --> foldr1 HOLogic.mk_prodT Ts') $ t) Ts'; |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
132 |
val projs = mk_proj (aux_lhs) Ts; |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
133 |
val proj_eqs = map2 (fn v => fn proj => (v, lambda arg proj)) vs projs; |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
134 |
val proj_defs = map2 (fn Free (name, _) => fn (_, rhs) => |
33205 | 135 |
((Binding.conceal (Binding.name name), NoSyn), |
33280 | 136 |
(apfst Binding.conceal Attrib.empty_binding, rhs))) vs proj_eqs; |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
137 |
val aux_eq' = Pattern.rewrite_term thy proj_eqs [] aux_eq; |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
138 |
fun prove_eqs aux_simp proj_defs lthy = |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
139 |
let |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
140 |
val proj_simps = map (snd o snd) proj_defs; |
31645 | 141 |
fun tac { context = ctxt, prems = _ } = |
31625 | 142 |
ALLGOALS (simp_tac (HOL_ss addsimps proj_simps)) |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
143 |
THEN ALLGOALS (EqSubst.eqsubst_tac ctxt [0] [aux_simp]) |
37136 | 144 |
THEN ALLGOALS (simp_tac (HOL_ss addsimps [@{thm fst_conv}, @{thm snd_conv}])); |
32970
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents:
32740
diff
changeset
|
145 |
in (map (fn prop => Skip_Proof.prove lthy [v] [] prop tac) eqs, lthy) end; |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
146 |
in |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
147 |
lthy |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
148 |
|> random_aux_primrec aux_eq' |
33766
c679f05600cd
adapted Local_Theory.define -- eliminated odd thm kind;
wenzelm
parents:
33671
diff
changeset
|
149 |
||>> fold_map Local_Theory.define proj_defs |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
150 |
|-> (fn (aux_simp, proj_defs) => prove_eqs aux_simp proj_defs) |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
151 |
end; |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
152 |
|
31868 | 153 |
fun random_aux_specification prfx name eqs lthy = |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
154 |
let |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
155 |
val vs = fold Term.add_free_names ((snd o strip_comb o fst o HOLogic.dest_eq |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
156 |
o HOLogic.dest_Trueprop o hd) eqs) []; |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
157 |
fun mk_proto_eq eq = |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
158 |
let |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
159 |
val (head $ t $ u, rhs) = (HOLogic.dest_eq o HOLogic.dest_Trueprop) eq; |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
160 |
in ((HOLogic.mk_Trueprop o HOLogic.mk_eq) (head, lambda t (lambda u rhs))) end; |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
161 |
val proto_eqs = map mk_proto_eq eqs; |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
162 |
fun prove_simps proto_simps lthy = |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
163 |
let |
31625 | 164 |
val ext_simps = map (fn thm => fun_cong OF [fun_cong OF [thm]]) proto_simps; |
42361 | 165 |
val tac = ALLGOALS (Proof_Context.fact_tac ext_simps); |
32970
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents:
32740
diff
changeset
|
166 |
in (map (fn prop => Skip_Proof.prove lthy vs [] prop (K tac)) eqs, lthy) end; |
33205 | 167 |
val b = Binding.conceal (Binding.qualify true prfx |
168 |
(Binding.qualify true name (Binding.name "simps"))); |
|
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
169 |
in |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
170 |
lthy |
31868 | 171 |
|> random_aux_primrec_multi (name ^ prfx) proto_eqs |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
172 |
|-> (fn proto_simps => prove_simps proto_simps) |
33671 | 173 |
|-> (fn simps => Local_Theory.note |
45592 | 174 |
((b, Code.add_default_eqn_attrib :: @{attributes [simp, nitpick_simp]}), simps)) |
31485
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
175 |
|> snd |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
176 |
end |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
177 |
|
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
178 |
|
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
179 |
(* constructing random instances on datatypes *) |
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
haftmann
parents:
31260
diff
changeset
|
180 |
|
31868 | 181 |
val random_auxN = "random_aux"; |
182 |
||
31603
fa30cd74d7d6
revised interpretation combinator for datatype constructions
haftmann
parents:
31595
diff
changeset
|
183 |
fun mk_random_aux_eqs thy descr vs tycos (names, auxnames) (Ts, Us) = |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
184 |
let |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
185 |
val mk_const = curry (Sign.mk_const thy); |
31603
fa30cd74d7d6
revised interpretation combinator for datatype constructions
haftmann
parents:
31595
diff
changeset
|
186 |
val random_auxsN = map (prefix (random_auxN ^ "_")) (names @ auxnames); |
fa30cd74d7d6
revised interpretation combinator for datatype constructions
haftmann
parents:
31595
diff
changeset
|
187 |
val rTs = Ts @ Us; |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
188 |
fun random_resultT T = @{typ Random.seed} |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
189 |
--> HOLogic.mk_prodT (termifyT T,@{typ Random.seed}); |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
190 |
val pTs = map random_resultT rTs; |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
191 |
fun sizeT T = @{typ code_numeral} --> @{typ code_numeral} --> T; |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
192 |
val random_auxT = sizeT o random_resultT; |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
193 |
val random_auxs = map2 (fn s => fn rT => Free (s, random_auxT rT)) |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
194 |
random_auxsN rTs; |
31950 | 195 |
fun mk_random_call T = (NONE, (HOLogic.mk_random T size', T)); |
31603
fa30cd74d7d6
revised interpretation combinator for datatype constructions
haftmann
parents:
31595
diff
changeset
|
196 |
fun mk_random_aux_call fTs (k, _) (tyco, Ts) = |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
197 |
let |
31623
b72c11302b39
quickcheck generators for datatypes with functions
haftmann
parents:
31611
diff
changeset
|
198 |
val T = Type (tyco, Ts); |
b72c11302b39
quickcheck generators for datatypes with functions
haftmann
parents:
31611
diff
changeset
|
199 |
fun mk_random_fun_lift [] t = t |
b72c11302b39
quickcheck generators for datatypes with functions
haftmann
parents:
31611
diff
changeset
|
200 |
| mk_random_fun_lift (fT :: fTs) t = |
b72c11302b39
quickcheck generators for datatypes with functions
haftmann
parents:
31611
diff
changeset
|
201 |
mk_const @{const_name random_fun_lift} [fTs ---> T, fT] $ |
b72c11302b39
quickcheck generators for datatypes with functions
haftmann
parents:
31611
diff
changeset
|
202 |
mk_random_fun_lift fTs t; |
31984 | 203 |
val t = mk_random_fun_lift fTs (nth random_auxs k $ size_pred $ size'); |
33968
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents:
33766
diff
changeset
|
204 |
val size = Option.map snd (Datatype_Aux.find_shortest_path descr k) |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
205 |
|> the_default 0; |
31623
b72c11302b39
quickcheck generators for datatypes with functions
haftmann
parents:
31611
diff
changeset
|
206 |
in (SOME size, (t, fTs ---> T)) end; |
33968
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents:
33766
diff
changeset
|
207 |
val tss = Datatype_Aux.interpret_construction descr vs |
31603
fa30cd74d7d6
revised interpretation combinator for datatype constructions
haftmann
parents:
31595
diff
changeset
|
208 |
{ atyp = mk_random_call, dtyp = mk_random_aux_call }; |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
209 |
fun mk_consexpr simpleT (c, xs) = |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
210 |
let |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
211 |
val (ks, simple_tTs) = split_list xs; |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
212 |
val T = termifyT simpleT; |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
213 |
val tTs = (map o apsnd) termifyT simple_tTs; |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
214 |
val is_rec = exists is_some ks; |
33029 | 215 |
val k = fold (fn NONE => I | SOME k => Integer.max k) ks 0; |
43329
84472e198515
tuned signature: Name.invent and Name.invent_names;
wenzelm
parents:
42361
diff
changeset
|
216 |
val vs = Name.invent_names Name.context "x" (map snd simple_tTs); |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
217 |
val tc = HOLogic.mk_return T @{typ Random.seed} |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
218 |
(HOLogic.mk_valtermify_app c vs simpleT); |
38543 | 219 |
val t = HOLogic.mk_ST |
220 |
(map2 (fn (t, _) => fn (v, T') => ((t, @{typ Random.seed}), SOME ((v, termifyT T')))) tTs vs) |
|
221 |
tc @{typ Random.seed} (SOME T, @{typ Random.seed}); |
|
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
222 |
val tk = if is_rec |
31950 | 223 |
then if k = 0 then size |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
224 |
else @{term "Quickcheck.beyond :: code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"} |
31950 | 225 |
$ HOLogic.mk_number @{typ code_numeral} k $ size |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
226 |
else @{term "1::code_numeral"} |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
227 |
in (is_rec, HOLogic.mk_prod (tk, t)) end; |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
228 |
fun sort_rec xs = |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
229 |
map_filter (fn (true, t) => SOME t | _ => NONE) xs |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
230 |
@ map_filter (fn (false, t) => SOME t | _ => NONE) xs; |
31603
fa30cd74d7d6
revised interpretation combinator for datatype constructions
haftmann
parents:
31595
diff
changeset
|
231 |
val gen_exprss = tss |
fa30cd74d7d6
revised interpretation combinator for datatype constructions
haftmann
parents:
31595
diff
changeset
|
232 |
|> (map o apfst) Type |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
233 |
|> map (fn (T, cs) => (T, (sort_rec o map (mk_consexpr T)) cs)); |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
234 |
fun mk_select (rT, xs) = |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
235 |
mk_const @{const_name Quickcheck.collapse} [@{typ "Random.seed"}, termifyT rT] |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
236 |
$ (mk_const @{const_name Random.select_weight} [random_resultT rT] |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
237 |
$ HOLogic.mk_list (HOLogic.mk_prodT (@{typ code_numeral}, random_resultT rT)) xs) |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
238 |
$ seed; |
31950 | 239 |
val auxs_lhss = map (fn t => t $ size $ size' $ seed) random_auxs; |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
240 |
val auxs_rhss = map mk_select gen_exprss; |
31868 | 241 |
in (random_auxs, auxs_lhss ~~ auxs_rhss) end; |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
242 |
|
38543 | 243 |
fun instantiate_random_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy = |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
244 |
let |
33968
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents:
33766
diff
changeset
|
245 |
val _ = Datatype_Aux.message config "Creating quickcheck generators ..."; |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
246 |
val mk_prop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq; |
33968
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents:
33766
diff
changeset
|
247 |
fun mk_size_arg k = case Datatype_Aux.find_shortest_path descr k |
31950 | 248 |
of SOME (_, l) => if l = 0 then size |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
249 |
else @{term "max :: code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"} |
31950 | 250 |
$ HOLogic.mk_number @{typ code_numeral} l $ size |
251 |
| NONE => size; |
|
31868 | 252 |
val (random_auxs, auxs_eqs) = (apsnd o map) mk_prop_eq |
31603
fa30cd74d7d6
revised interpretation combinator for datatype constructions
haftmann
parents:
31595
diff
changeset
|
253 |
(mk_random_aux_eqs thy descr vs tycos (names, auxnames) (Ts, Us)); |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
254 |
val random_defs = map_index (fn (k, T) => mk_prop_eq |
31950 | 255 |
(HOLogic.mk_random T size, nth random_auxs k $ mk_size_arg k $ size)) Ts; |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
256 |
in |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
257 |
thy |
38348
cf7b2121ad9d
moved instantiation target formally to class_target.ML
haftmann
parents:
37744
diff
changeset
|
258 |
|> Class.instantiation (tycos, vs, @{sort random}) |
31868 | 259 |
|> random_aux_specification prfx random_auxN auxs_eqs |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
260 |
|> `(fn lthy => map (Syntax.check_term lthy) random_defs) |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
261 |
|-> (fn random_defs' => fold_map (fn random_def => |
33280 | 262 |
Specification.definition (NONE, (apfst Binding.conceal |
33205 | 263 |
Attrib.empty_binding, random_def))) random_defs') |
31595
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
264 |
|> snd |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
265 |
|> Class.prove_instantiation_exit (K (Class.intro_classes_tac [])) |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
266 |
end; |
bd2f7211a420
first running version of qc generators for datatypes
haftmann
parents:
31485
diff
changeset
|
267 |
|
31950 | 268 |
(** building and compiling generator expressions **) |
269 |
||
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40911
diff
changeset
|
270 |
(* FIXME just one data slot (record) per program unit *) |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40911
diff
changeset
|
271 |
|
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40911
diff
changeset
|
272 |
structure Counterexample = Proof_Data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40911
diff
changeset
|
273 |
( |
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
274 |
type T = unit -> int -> int -> int * int -> term list option * (int * int) |
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40911
diff
changeset
|
275 |
(* FIXME avoid user error with non-user text *) |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
276 |
fun init _ () = error "Counterexample" |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
277 |
); |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
278 |
val put_counterexample = Counterexample.put; |
31950 | 279 |
|
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40911
diff
changeset
|
280 |
structure Counterexample_Report = Proof_Data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40911
diff
changeset
|
281 |
( |
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
282 |
type T = unit -> int -> int -> seed -> (term list option * (bool list * bool)) * seed |
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40911
diff
changeset
|
283 |
(* FIXME avoid user error with non-user text *) |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
284 |
fun init _ () = error "Counterexample_Report" |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
285 |
); |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
286 |
val put_counterexample_report = Counterexample_Report.put; |
35378
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
287 |
|
31950 | 288 |
val target = "Quickcheck"; |
289 |
||
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
290 |
fun mk_generator_expr ctxt (t, eval_terms) = |
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
291 |
let |
42361 | 292 |
val thy = Proof_Context.theory_of ctxt |
44241 | 293 |
val prop = fold_rev absfree (Term.add_frees t []) t |
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
294 |
val Ts = (map snd o fst o strip_abs) prop |
31950 | 295 |
val bound_max = length Ts - 1; |
296 |
val bounds = map_index (fn (i, ty) => |
|
297 |
(2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) Ts; |
|
298 |
val result = list_comb (prop, map (fn (i, _, _, _) => Bound i) bounds); |
|
299 |
val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) bounds); |
|
45718
8979b2463fc8
quickcheck random can also find potential counterexamples;
bulwahn
parents:
45592
diff
changeset
|
300 |
val check = Quickcheck_Common.mk_safe_if ctxt (result, @{term "None :: term list option"}, |
8979b2463fc8
quickcheck random can also find potential counterexamples;
bulwahn
parents:
45592
diff
changeset
|
301 |
@{term "Some :: term list => term list option"} $ terms) |
31950 | 302 |
val return = @{term "Pair :: term list option => Random.seed => term list option * Random.seed"}; |
303 |
fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT); |
|
304 |
fun mk_termtyp T = HOLogic.mk_prodT (T, @{typ "unit => term"}); |
|
305 |
fun mk_scomp T1 T2 sT f g = Const (@{const_name scomp}, |
|
306 |
liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g; |
|
307 |
fun mk_split T = Sign.mk_const thy |
|
37591 | 308 |
(@{const_name prod_case}, [T, @{typ "unit => term"}, liftT @{typ "term list option"} @{typ Random.seed}]); |
31950 | 309 |
fun mk_scomp_split T t t' = |
310 |
mk_scomp (mk_termtyp T) @{typ "term list option"} @{typ Random.seed} t |
|
311 |
(mk_split T $ Abs ("", T, Abs ("", @{typ "unit => term"}, t'))); |
|
312 |
fun mk_bindclause (_, _, i, T) = mk_scomp_split T |
|
313 |
(Sign.mk_const thy (@{const_name Quickcheck.random}, [T]) $ Bound i); |
|
314 |
in Abs ("n", @{typ code_numeral}, fold_rev mk_bindclause bounds (return $ check)) end; |
|
315 |
||
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
316 |
fun mk_reporting_generator_expr ctxt (t, eval_terms) = |
35378
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
317 |
let |
42361 | 318 |
val thy = Proof_Context.theory_of ctxt |
44241 | 319 |
val prop = fold_rev absfree (Term.add_frees t []) t |
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
320 |
val Ts = (map snd o fst o strip_abs) prop |
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
321 |
val bound_max = length Ts - 1 |
35378
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
322 |
val bounds = map_index (fn (i, ty) => |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
323 |
(2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) Ts; |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
324 |
val prop' = betapplys (prop, map (fn (i, _, _, _) => Bound i) bounds); |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
325 |
val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) bounds) |
42195
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42159
diff
changeset
|
326 |
val (assms, concl) = Quickcheck_Common.strip_imp prop' |
35378
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
327 |
val return = |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
328 |
@{term "Pair :: term list option * (bool list * bool) => Random.seed => (term list option * (bool list * bool)) * Random.seed"}; |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
329 |
fun mk_assms_report i = |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
330 |
HOLogic.mk_prod (@{term "None :: term list option"}, |
38553
56965d8e4e11
use HOLogic.boolT and @{typ bool} more pervasively
haftmann
parents:
38549
diff
changeset
|
331 |
HOLogic.mk_prod (HOLogic.mk_list HOLogic.boolT |
56965d8e4e11
use HOLogic.boolT and @{typ bool} more pervasively
haftmann
parents:
38549
diff
changeset
|
332 |
(replicate i @{term True} @ replicate (length assms - i) @{term False}), |
56965d8e4e11
use HOLogic.boolT and @{typ bool} more pervasively
haftmann
parents:
38549
diff
changeset
|
333 |
@{term False})) |
35378
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
334 |
fun mk_concl_report b = |
38553
56965d8e4e11
use HOLogic.boolT and @{typ bool} more pervasively
haftmann
parents:
38549
diff
changeset
|
335 |
HOLogic.mk_prod (HOLogic.mk_list HOLogic.boolT (replicate (length assms) @{term True}), |
35378
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
336 |
if b then @{term True} else @{term False}) |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
337 |
val If = |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
338 |
@{term "If :: bool => term list option * (bool list * bool) => term list option * (bool list * bool) => term list option * (bool list * bool)"} |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
339 |
val concl_check = If $ concl $ |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
340 |
HOLogic.mk_prod (@{term "None :: term list option"}, mk_concl_report true) $ |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
341 |
HOLogic.mk_prod (@{term "Some :: term list => term list option"} $ terms, mk_concl_report false) |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
342 |
val check = fold_rev (fn (i, assm) => fn t => If $ assm $ t $ mk_assms_report i) |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
343 |
(map_index I assms) concl_check |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
344 |
fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT); |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
345 |
fun mk_termtyp T = HOLogic.mk_prodT (T, @{typ "unit => term"}); |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
346 |
fun mk_scomp T1 T2 sT f g = Const (@{const_name scomp}, |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
347 |
liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g; |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
348 |
fun mk_split T = Sign.mk_const thy |
37591 | 349 |
(@{const_name prod_case}, [T, @{typ "unit => term"}, |
35378
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
350 |
liftT @{typ "term list option * (bool list * bool)"} @{typ Random.seed}]); |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
351 |
fun mk_scomp_split T t t' = |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
352 |
mk_scomp (mk_termtyp T) @{typ "term list option * (bool list * bool)"} @{typ Random.seed} t |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
353 |
(mk_split T $ Abs ("", T, Abs ("", @{typ "unit => term"}, t'))); |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
354 |
fun mk_bindclause (_, _, i, T) = mk_scomp_split T |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
355 |
(Sign.mk_const thy (@{const_name Quickcheck.random}, [T]) $ Bound i); |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
356 |
in |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
357 |
Abs ("n", @{typ code_numeral}, fold_rev mk_bindclause bounds (return $ check)) |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
358 |
end |
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
359 |
|
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
360 |
val mk_parametric_generator_expr = Quickcheck_Common.gen_mk_parametric_generator_expr |
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
361 |
((mk_generator_expr, |
44241 | 362 |
absdummy @{typ code_numeral} @{term "Pair None :: Random.seed => term list option * Random.seed"}), |
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
363 |
@{typ "code_numeral => Random.seed => term list option * Random.seed"}) |
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
364 |
|
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
365 |
val mk_parametric_reporting_generator_expr = Quickcheck_Common.gen_mk_parametric_generator_expr |
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
366 |
((mk_reporting_generator_expr, |
44241 | 367 |
absdummy @{typ code_numeral} |
368 |
@{term "Pair (None, ([], False)) :: Random.seed => (term list option * (bool list * bool)) * Random.seed"}), |
|
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
369 |
@{typ "code_numeral => Random.seed => (term list option * (bool list * bool)) * Random.seed"}) |
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
370 |
|
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
371 |
|
40911
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
372 |
(* single quickcheck report *) |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
373 |
|
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
374 |
datatype single_report = Run of bool list * bool | MatchExc |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
375 |
|
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
376 |
fun collect_single_report single_report |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
377 |
(Quickcheck.Report {iterations = iterations, raised_match_errors = raised_match_errors, |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
378 |
satisfied_assms = satisfied_assms, positive_concl_tests = positive_concl_tests}) = |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
379 |
case single_report |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
380 |
of MatchExc => |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
381 |
Quickcheck.Report {iterations = iterations + 1, raised_match_errors = raised_match_errors + 1, |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
382 |
satisfied_assms = satisfied_assms, positive_concl_tests = positive_concl_tests} |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
383 |
| Run (assms, concl) => |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
384 |
Quickcheck.Report {iterations = iterations + 1, raised_match_errors = raised_match_errors, |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
385 |
satisfied_assms = |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
386 |
map2 (fn b => fn s => if b then s + 1 else s) assms |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
387 |
(if null satisfied_assms then replicate (length assms) 0 else satisfied_assms), |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
388 |
positive_concl_tests = if concl then positive_concl_tests + 1 else positive_concl_tests} |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
389 |
|
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
390 |
val empty_report = Quickcheck.Report { iterations = 0, raised_match_errors = 0, |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
391 |
satisfied_assms = [], positive_concl_tests = 0 } |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
392 |
|
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
393 |
fun compile_generator_expr ctxt ts = |
31950 | 394 |
let |
42361 | 395 |
val thy = Proof_Context.theory_of ctxt |
40911
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
396 |
val iterations = Config.get ctxt Quickcheck.iterations |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
397 |
in |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
398 |
if Config.get ctxt Quickcheck.report then |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
399 |
let |
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
400 |
val t' = mk_parametric_reporting_generator_expr ctxt ts; |
40911
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
401 |
val compile = Code_Runtime.dynamic_value_strict |
41923
f05fc0711bc7
renaming signatures and structures; correcting header
bulwahn
parents:
41921
diff
changeset
|
402 |
(Counterexample_Report.get, put_counterexample_report, "Random_Generators.put_counterexample_report") |
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
403 |
thy (SOME target) (fn proc => fn g => fn c => fn s => g c s #>> (apfst o Option.map o map) proc) t' []; |
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
404 |
fun single_tester c s = compile c s |> Random_Engine.run |
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
405 |
fun iterate_and_collect (card, size) 0 report = (NONE, report) |
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
406 |
| iterate_and_collect (card, size) j report = |
40911
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
407 |
let |
45719
39c52a820f6e
removing exception handling now that is caught at some other point;
bulwahn
parents:
45718
diff
changeset
|
408 |
val (test_result, single_report) = apsnd Run (single_tester card size) |
40911
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
409 |
val report = collect_single_report single_report report |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
410 |
in |
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
411 |
case test_result of NONE => iterate_and_collect (card, size) (j - 1) report |
40911
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
412 |
| SOME q => (SOME q, report) |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
413 |
end |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
414 |
in |
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
415 |
fn [card, size] => apsnd SOME (iterate_and_collect (card, size) iterations empty_report) |
40911
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
416 |
end |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
417 |
else |
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
418 |
let |
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
419 |
val t' = mk_parametric_generator_expr ctxt ts; |
40911
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
420 |
val compile = Code_Runtime.dynamic_value_strict |
41923
f05fc0711bc7
renaming signatures and structures; correcting header
bulwahn
parents:
41921
diff
changeset
|
421 |
(Counterexample.get, put_counterexample, "Random_Generators.put_counterexample") |
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
422 |
thy (SOME target) (fn proc => fn g => fn c => fn s => g c s #>> (Option.map o map) proc) t' []; |
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
423 |
fun single_tester c s = compile c s |> Random_Engine.run |
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
424 |
fun iterate (card, size) 0 = NONE |
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
425 |
| iterate (card, size) j = |
45719
39c52a820f6e
removing exception handling now that is caught at some other point;
bulwahn
parents:
45718
diff
changeset
|
426 |
case single_tester card size of NONE => iterate (card, size) (j - 1) | SOME q => SOME q |
40911
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
427 |
in |
42159
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
428 |
fn [card, size] => (rpair NONE (iterate (card, size) iterations)) |
40911
7febf76e0a69
moving iteration of tests to the testers in quickcheck
bulwahn
parents:
40644
diff
changeset
|
429 |
end |
35378
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35166
diff
changeset
|
430 |
end; |
31950 | 431 |
|
45420
d17556b9a89b
more precise messages with the tester's name in quickcheck; tuned
bulwahn
parents:
45159
diff
changeset
|
432 |
val test_goals = Quickcheck_Common.generator_test_goal_terms ("random", compile_generator_expr); |
43875
485d2ad43528
adding random, exhaustive and SML quickcheck as testers
bulwahn
parents:
43333
diff
changeset
|
433 |
|
31260 | 434 |
(** setup **) |
435 |
||
43878
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
bulwahn
parents:
43877
diff
changeset
|
436 |
val active = Attrib.setup_config_bool @{binding quickcheck_random_active} (K false); |
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
bulwahn
parents:
43877
diff
changeset
|
437 |
|
38393 | 438 |
val setup = |
42229
1491b7209e76
generalizing ensure_sort_datatype for bounded_forall instances
bulwahn
parents:
42195
diff
changeset
|
439 |
Datatype.interpretation (Quickcheck_Common.ensure_sort_datatype |
1491b7209e76
generalizing ensure_sort_datatype for bounded_forall instances
bulwahn
parents:
42195
diff
changeset
|
440 |
(((@{sort typerep}, @{sort term_of}), @{sort random}), instantiate_random_datatype)) |
43878
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
bulwahn
parents:
43877
diff
changeset
|
441 |
#> Context.theory_map (Quickcheck.add_tester ("random", (active, test_goals))); |
31260 | 442 |
|
443 |
end; |