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