| author | wenzelm | 
| Sat, 04 Mar 2023 12:14:20 +0100 | |
| changeset 77500 | bbb78dba6f68 | 
| parent 74383 | 107941e8fa01 | 
| child 77879 | dd222e2af01a | 
| 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  | 
|
| 62979 | 10  | 
  val random_fun: typ -> typ -> ('a -> 'a -> bool) -> ('a -> term) ->
 | 
11  | 
    (seed -> ('b * (unit -> term)) * seed) -> (seed -> seed * seed) ->
 | 
|
12  | 
    seed -> (('a -> 'b) * (unit -> term)) * seed
 | 
|
13  | 
val compile_generator_expr: Proof.context -> (term * term list) list -> bool -> int list ->  | 
|
14  | 
(bool * term list) option * Quickcheck.report option  | 
|
15  | 
val put_counterexample: (unit -> Code_Numeral.natural -> bool -> Code_Numeral.natural ->  | 
|
16  | 
seed -> (bool * term list) option * seed) -> Proof.context -> Proof.context  | 
|
17  | 
val put_counterexample_report: (unit -> Code_Numeral.natural -> bool ->  | 
|
18  | 
Code_Numeral.natural -> seed -> ((bool * term list) option * (bool list * bool)) * seed) ->  | 
|
19  | 
Proof.context -> Proof.context  | 
|
| 
58112
 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 
blanchet 
parents: 
55757 
diff
changeset
 | 
20  | 
val instantiate_random_datatype : Old_Datatype_Aux.config -> Old_Datatype_Aux.descr ->  | 
| 62979 | 21  | 
(string * sort) list -> string list -> string -> string list * string list ->  | 
22  | 
typ list * typ list -> theory -> theory  | 
|
| 31260 | 23  | 
end;  | 
24  | 
||
| 
41923
 
f05fc0711bc7
renaming signatures and structures; correcting header
 
bulwahn 
parents: 
41921 
diff
changeset
 | 
25  | 
structure Random_Generators : RANDOM_GENERATORS =  | 
| 31260 | 26  | 
struct  | 
27  | 
||
| 31950 | 28  | 
(** abstract syntax **)  | 
| 31260 | 29  | 
|
| 67149 | 30  | 
fun termifyT T = HOLogic.mk_prodT (T, \<^typ>\<open>unit \<Rightarrow> term\<close>)  | 
31  | 
val size = \<^term>\<open>i::natural\<close>;  | 
|
32  | 
val size_pred = \<^term>\<open>(i::natural) - 1\<close>;  | 
|
33  | 
val size' = \<^term>\<open>j::natural\<close>;  | 
|
34  | 
val seed = \<^term>\<open>s::Random.seed\<close>;  | 
|
| 31260 | 35  | 
|
| 67149 | 36  | 
val resultT = \<^typ>\<open>(bool \<times> term list) option\<close>;  | 
| 
45721
 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 
bulwahn 
parents: 
45719 
diff
changeset
 | 
37  | 
|
| 62979 | 38  | 
|
39  | 
(** typ "'a \<Rightarrow> 'b" **)  | 
|
| 31260 | 40  | 
|
41  | 
type seed = Random_Engine.seed;  | 
|
42  | 
||
| 
31603
 
fa30cd74d7d6
revised interpretation combinator for datatype constructions
 
haftmann 
parents: 
31595 
diff
changeset
 | 
43  | 
fun random_fun T1 T2 eq term_of random random_split seed =  | 
| 31260 | 44  | 
let  | 
| 67149 | 45  | 
val fun_upd = Const (\<^const_name>\<open>fun_upd\<close>, (T1 --> T2) --> T1 --> T2 --> T1 --> T2);  | 
| 50046 | 46  | 
val ((_, t2), seed') = random seed;  | 
| 32344 | 47  | 
val (seed'', seed''') = random_split seed';  | 
| 
31933
 
cd6511035315
proper closures -- imperative programming considered harmful...
 
haftmann 
parents: 
31902 
diff
changeset
 | 
48  | 
|
| 32740 | 49  | 
    val state = Unsynchronized.ref (seed'', [], fn () => Abs ("x", T1, t2 ()));
 | 
| 31260 | 50  | 
fun random_fun' x =  | 
51  | 
let  | 
|
52  | 
val (seed, fun_map, f_t) = ! state;  | 
|
53  | 
in case AList.lookup (uncurry eq) fun_map x  | 
|
54  | 
of SOME y => y  | 
|
55  | 
| NONE => let  | 
|
56  | 
val t1 = term_of x;  | 
|
57  | 
val ((y, t2), seed') = random seed;  | 
|
58  | 
val fun_map' = (x, y) :: fun_map;  | 
|
| 
31933
 
cd6511035315
proper closures -- imperative programming considered harmful...
 
haftmann 
parents: 
31902 
diff
changeset
 | 
59  | 
val f_t' = fn () => fun_upd $ f_t () $ t1 $ t2 ();  | 
| 31260 | 60  | 
val _ = state := (seed', fun_map', f_t');  | 
61  | 
in y end  | 
|
62  | 
end;  | 
|
| 
31933
 
cd6511035315
proper closures -- imperative programming considered harmful...
 
haftmann 
parents: 
31902 
diff
changeset
 | 
63  | 
fun term_fun' () = #3 (! state) ();  | 
| 32344 | 64  | 
in ((random_fun', term_fun'), seed''') end;  | 
| 31260 | 65  | 
|
| 
45762
 
daf57640d4df
the reporting random testing also returns if the counterexample is genuine or potentially spurious
 
bulwahn 
parents: 
45761 
diff
changeset
 | 
66  | 
|
| 31260 | 67  | 
(** datatypes **)  | 
68  | 
||
| 
31485
 
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
 
haftmann 
parents: 
31260 
diff
changeset
 | 
69  | 
(* definitional scheme for random instances on datatypes *)  | 
| 
 
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
 
haftmann 
parents: 
31260 
diff
changeset
 | 
70  | 
|
| 
31611
 
a577f77af93f
explicit instantiation yields considerable speedup
 
haftmann 
parents: 
31609 
diff
changeset
 | 
71  | 
local  | 
| 
 
a577f77af93f
explicit instantiation yields considerable speedup
 
haftmann 
parents: 
31609 
diff
changeset
 | 
72  | 
|
| 
 
a577f77af93f
explicit instantiation yields considerable speedup
 
haftmann 
parents: 
31609 
diff
changeset
 | 
73  | 
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
 | 
74  | 
val lhs = eq |> Thm.dest_arg1;  | 
| 
 
a577f77af93f
explicit instantiation yields considerable speedup
 
haftmann 
parents: 
31609 
diff
changeset
 | 
75  | 
val pt_random_aux = lhs |> Thm.dest_fun;  | 
| 
 
a577f77af93f
explicit instantiation yields considerable speedup
 
haftmann 
parents: 
31609 
diff
changeset
 | 
76  | 
val pt_rhs = eq |> Thm.dest_arg |> Thm.dest_fun;  | 
| 
60642
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60003 
diff
changeset
 | 
77  | 
val a_v =  | 
| 70159 | 78  | 
pt_random_aux |> Thm.ctyp_of_cterm |> Thm.dest_ctyp1  | 
| 
60642
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60003 
diff
changeset
 | 
79  | 
|> Thm.typ_of |> dest_TVar;  | 
| 
31611
 
a577f77af93f
explicit instantiation yields considerable speedup
 
haftmann 
parents: 
31609 
diff
changeset
 | 
80  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
81  | 
val rew_thms = map mk_meta_eq [@{thm natural_zero_minus_one},
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
82  | 
  @{thm Suc_natural_minus_one}, @{thm select_weight_cons_zero}, @{thm beyond_zero}];
 | 
| 
31611
 
a577f77af93f
explicit instantiation yields considerable speedup
 
haftmann 
parents: 
31609 
diff
changeset
 | 
83  | 
val rew_ts = map (Logic.dest_equals o Thm.prop_of) rew_thms;  | 
| 69593 | 84  | 
val rew_ss = simpset_of (put_simpset HOL_ss \<^context> addsimps rew_thms);  | 
| 
31611
 
a577f77af93f
explicit instantiation yields considerable speedup
 
haftmann 
parents: 
31609 
diff
changeset
 | 
85  | 
|
| 
 
a577f77af93f
explicit instantiation yields considerable speedup
 
haftmann 
parents: 
31609 
diff
changeset
 | 
86  | 
in  | 
| 
31485
 
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
 
haftmann 
parents: 
31260 
diff
changeset
 | 
87  | 
|
| 
 
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
 
haftmann 
parents: 
31260 
diff
changeset
 | 
88  | 
fun random_aux_primrec eq lthy =  | 
| 
 
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
 
haftmann 
parents: 
31260 
diff
changeset
 | 
89  | 
let  | 
| 42361 | 90  | 
val thy = Proof_Context.theory_of lthy;  | 
| 
31611
 
a577f77af93f
explicit instantiation yields considerable speedup
 
haftmann 
parents: 
31609 
diff
changeset
 | 
91  | 
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
 | 
92  | 
(HOLogic.dest_eq o HOLogic.dest_Trueprop) eq;  | 
| 
31485
 
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
 
haftmann 
parents: 
31260 
diff
changeset
 | 
93  | 
val Type (_, [_, iT]) = T;  | 
| 59637 | 94  | 
val icT = Thm.ctyp_of lthy iT;  | 
| 74282 | 95  | 
val inst = Thm.instantiate_cterm (TVars.make [(a_v, icT)], Vars.empty);  | 
| 
31485
 
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
 
haftmann 
parents: 
31260 
diff
changeset
 | 
96  | 
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
 | 
97  | 
val t_rhs = lambda t_k proto_t_rhs;  | 
| 67149 | 98  | 
val eqs0 = [subst_v \<^term>\<open>0::natural\<close> eq,  | 
| 74383 | 99  | 
subst_v \<^Const>\<open>Code_Numeral.Suc for t_k\<close> eq];  | 
| 
31611
 
a577f77af93f
explicit instantiation yields considerable speedup
 
haftmann 
parents: 
31609 
diff
changeset
 | 
100  | 
val eqs1 = map (Pattern.rewrite_term thy rew_ts []) eqs0;  | 
| 59617 | 101  | 
val ((_, (_, eqs2)), lthy') = lthy  | 
| 60003 | 102  | 
|> BNF_LFP_Compat.primrec_simple  | 
| 59859 | 103  | 
[((Binding.concealed (Binding.name random_aux), T), NoSyn)] eqs1;  | 
| 
60642
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60003 
diff
changeset
 | 
104  | 
val cT_random_aux = inst pt_random_aux |> Thm.term_of |> dest_Var;  | 
| 
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60003 
diff
changeset
 | 
105  | 
val cT_rhs = inst pt_rhs |> Thm.term_of |> dest_Var;  | 
| 
31485
 
259a3c90016e
added infrastructure for definitorial construction of generators for datatypes
 
haftmann 
parents: 
31260 
diff
changeset
 | 
106  | 
    val rule = @{thm random_aux_rec}
 | 
| 59617 | 107  | 
|> Drule.instantiate_normalize  | 
| 74282 | 108  | 
(TVars.make [(a_v, icT)],  | 
109  | 
Vars.make [(cT_random_aux, Thm.cterm_of lthy' t_random_aux),  | 
|
| 59637 | 110  | 
(cT_rhs, Thm.cterm_of lthy' t_rhs)]);  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51551 
diff
changeset
 | 
111  | 
fun tac ctxt =  | 
| 60752 | 112  | 
ALLGOALS (resolve_tac ctxt [rule])  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51551 
diff
changeset
 | 
113  | 
THEN ALLGOALS (simp_tac (put_simpset rew_ss ctxt))  | 
| 59617 | 114  | 
THEN (ALLGOALS (Proof_Context.fact_tac ctxt eqs2));  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51551 
diff
changeset
 | 
115  | 
val simp = Goal.prove_sorry lthy' [v] [] eq (tac o #context);  | 
| 
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  | 
| 58229 | 123  | 
|>> single  | 
| 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 _ :: _)) =  | 
| 67149 | 136  | 
Const (\<^const_name>\<open>fst\<close>, foldr1 HOLogic.mk_prodT Ts --> T) $ t  | 
137  | 
:: mk_proj (Const (\<^const_name>\<open>snd\<close>,  | 
|
| 
31485
 
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) =>  | 
| 59859 | 142  | 
((Binding.concealed (Binding.name name), NoSyn),  | 
| 63352 | 143  | 
(apfst Binding.concealed Binding.empty_atts, 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 = _ } =
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51551 
diff
changeset
 | 
149  | 
ALLGOALS (simp_tac (put_simpset HOL_ss ctxt 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])  | 
| 58229 | 151  | 
              THEN ALLGOALS (simp_tac (put_simpset HOL_ss ctxt addsimps @{thms fst_conv snd_conv}));
 | 
| 51551 | 152  | 
in (map (fn prop => Goal.prove_sorry 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  | 
| 58229 | 157  | 
|-> uncurry prove_eqs  | 
| 
31485
 
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;  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
51717 
diff
changeset
 | 
172  | 
val tac = ALLGOALS (Proof_Context.fact_tac lthy ext_simps);  | 
| 51551 | 173  | 
in (map (fn prop => Goal.prove_sorry lthy vs [] prop (K tac)) eqs, lthy) end;  | 
| 59859 | 174  | 
val b = Binding.concealed (Binding.qualify true prfx  | 
| 33205 | 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  | 
| 58229 | 179  | 
|-> prove_simps  | 
| 33671 | 180  | 
|-> (fn simps => Local_Theory.note  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
63352 
diff
changeset
 | 
181  | 
          ((b, @{attributes [simp, nitpick_simp]}), simps))
 | 
| 
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
63352 
diff
changeset
 | 
182  | 
|-> (fn (_, thms) => Code.declare_default_eqns (map (rpair true) thms))  | 
| 
31485
 
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  | 
||
| 50046 | 190  | 
fun mk_random_aux_eqs thy descr vs (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;  | 
| 67149 | 195  | 
fun random_resultT T = \<^typ>\<open>Random.seed\<close>  | 
196  | 
--> HOLogic.mk_prodT (termifyT T,\<^typ>\<open>Random.seed\<close>);  | 
|
197  | 
fun sizeT T = \<^typ>\<open>natural\<close> --> \<^typ>\<open>natural\<close> --> T;  | 
|
| 
31595
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
198  | 
val random_auxT = sizeT o random_resultT;  | 
| 
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
199  | 
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
 | 
200  | 
random_auxsN rTs;  | 
| 31950 | 201  | 
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
 | 
202  | 
fun mk_random_aux_call fTs (k, _) (tyco, Ts) =  | 
| 
31595
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
203  | 
let  | 
| 
31623
 
b72c11302b39
quickcheck generators for datatypes with functions
 
haftmann 
parents: 
31611 
diff
changeset
 | 
204  | 
val T = Type (tyco, Ts);  | 
| 
 
b72c11302b39
quickcheck generators for datatypes with functions
 
haftmann 
parents: 
31611 
diff
changeset
 | 
205  | 
fun mk_random_fun_lift [] t = t  | 
| 
 
b72c11302b39
quickcheck generators for datatypes with functions
 
haftmann 
parents: 
31611 
diff
changeset
 | 
206  | 
| mk_random_fun_lift (fT :: fTs) t =  | 
| 67149 | 207  | 
mk_const \<^const_name>\<open>random_fun_lift\<close> [fTs ---> T, fT] $  | 
| 
31623
 
b72c11302b39
quickcheck generators for datatypes with functions
 
haftmann 
parents: 
31611 
diff
changeset
 | 
208  | 
mk_random_fun_lift fTs t;  | 
| 31984 | 209  | 
val t = mk_random_fun_lift fTs (nth random_auxs k $ size_pred $ size');  | 
| 
58112
 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 
blanchet 
parents: 
55757 
diff
changeset
 | 
210  | 
val size = Option.map snd (Old_Datatype_Aux.find_shortest_path descr k)  | 
| 
31595
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
211  | 
|> the_default 0;  | 
| 
31623
 
b72c11302b39
quickcheck generators for datatypes with functions
 
haftmann 
parents: 
31611 
diff
changeset
 | 
212  | 
in (SOME size, (t, fTs ---> T)) end;  | 
| 
58112
 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 
blanchet 
parents: 
55757 
diff
changeset
 | 
213  | 
val tss = Old_Datatype_Aux.interpret_construction descr vs  | 
| 
31603
 
fa30cd74d7d6
revised interpretation combinator for datatype constructions
 
haftmann 
parents: 
31595 
diff
changeset
 | 
214  | 
      { atyp = mk_random_call, dtyp = mk_random_aux_call };
 | 
| 
31595
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
215  | 
fun mk_consexpr simpleT (c, xs) =  | 
| 
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
216  | 
let  | 
| 
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
217  | 
val (ks, simple_tTs) = split_list xs;  | 
| 
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
218  | 
val T = termifyT simpleT;  | 
| 
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
219  | 
val tTs = (map o apsnd) termifyT simple_tTs;  | 
| 
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
220  | 
val is_rec = exists is_some ks;  | 
| 33029 | 221  | 
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
 | 
222  | 
val vs = Name.invent_names Name.context "x" (map snd simple_tTs);  | 
| 67149 | 223  | 
val tc = HOLogic.mk_return T \<^typ>\<open>Random.seed\<close>  | 
| 
31595
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
224  | 
(HOLogic.mk_valtermify_app c vs simpleT);  | 
| 38543 | 225  | 
val t = HOLogic.mk_ST  | 
| 67149 | 226  | 
(map2 (fn (t, _) => fn (v, T') => ((t, \<^typ>\<open>Random.seed\<close>), SOME ((v, termifyT T')))) tTs vs)  | 
227  | 
tc \<^typ>\<open>Random.seed\<close> (SOME T, \<^typ>\<open>Random.seed\<close>);  | 
|
| 
31595
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
228  | 
val tk = if is_rec  | 
| 31950 | 229  | 
then if k = 0 then size  | 
| 67149 | 230  | 
else \<^term>\<open>Quickcheck_Random.beyond :: natural \<Rightarrow> natural \<Rightarrow> natural\<close>  | 
231  | 
$ HOLogic.mk_number \<^typ>\<open>natural\<close> k $ size  | 
|
232  | 
else \<^term>\<open>1::natural\<close>  | 
|
| 
31595
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
233  | 
in (is_rec, HOLogic.mk_prod (tk, t)) end;  | 
| 
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
234  | 
fun sort_rec xs =  | 
| 
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
235  | 
map_filter (fn (true, t) => SOME t | _ => NONE) xs  | 
| 
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
236  | 
@ map_filter (fn (false, t) => SOME t | _ => NONE) xs;  | 
| 
31603
 
fa30cd74d7d6
revised interpretation combinator for datatype constructions
 
haftmann 
parents: 
31595 
diff
changeset
 | 
237  | 
val gen_exprss = tss  | 
| 
 
fa30cd74d7d6
revised interpretation combinator for datatype constructions
 
haftmann 
parents: 
31595 
diff
changeset
 | 
238  | 
|> (map o apfst) Type  | 
| 
31595
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
239  | 
|> 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
 | 
240  | 
fun mk_select (rT, xs) =  | 
| 67149 | 241  | 
mk_const \<^const_name>\<open>Quickcheck_Random.collapse\<close> [\<^typ>\<open>Random.seed\<close>, termifyT rT]  | 
242  | 
$ (mk_const \<^const_name>\<open>Random.select_weight\<close> [random_resultT rT]  | 
|
243  | 
$ HOLogic.mk_list (HOLogic.mk_prodT (\<^typ>\<open>natural\<close>, random_resultT rT)) xs)  | 
|
| 
31595
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
244  | 
$ seed;  | 
| 31950 | 245  | 
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
 | 
246  | 
val auxs_rhss = map mk_select gen_exprss;  | 
| 31868 | 247  | 
in (random_auxs, auxs_lhss ~~ auxs_rhss) end;  | 
| 
31595
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
248  | 
|
| 38543 | 249  | 
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
 | 
250  | 
let  | 
| 
58112
 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 
blanchet 
parents: 
55757 
diff
changeset
 | 
251  | 
val _ = Old_Datatype_Aux.message config "Creating quickcheck generators ...";  | 
| 
31595
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
252  | 
val mk_prop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq;  | 
| 
58112
 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 
blanchet 
parents: 
55757 
diff
changeset
 | 
253  | 
fun mk_size_arg k = case Old_Datatype_Aux.find_shortest_path descr k  | 
| 31950 | 254  | 
of SOME (_, l) => if l = 0 then size  | 
| 67149 | 255  | 
else \<^term>\<open>max :: natural \<Rightarrow> natural \<Rightarrow> natural\<close>  | 
256  | 
$ HOLogic.mk_number \<^typ>\<open>natural\<close> l $ size  | 
|
| 31950 | 257  | 
| NONE => size;  | 
| 31868 | 258  | 
val (random_auxs, auxs_eqs) = (apsnd o map) mk_prop_eq  | 
| 50046 | 259  | 
(mk_random_aux_eqs thy descr vs (names, auxnames) (Ts, Us));  | 
| 
31595
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
260  | 
val random_defs = map_index (fn (k, T) => mk_prop_eq  | 
| 31950 | 261  | 
(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
 | 
262  | 
in  | 
| 
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
263  | 
thy  | 
| 67149 | 264  | 
|> Class.instantiation (tycos, vs, \<^sort>\<open>random\<close>)  | 
| 31868 | 265  | 
|> random_aux_specification prfx random_auxN auxs_eqs  | 
| 
31595
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
266  | 
|> `(fn lthy => map (Syntax.check_term lthy) random_defs)  | 
| 
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
267  | 
|-> (fn random_defs' => fold_map (fn random_def =>  | 
| 63180 | 268  | 
Specification.definition NONE [] []  | 
| 
63064
 
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
 
wenzelm 
parents: 
62979 
diff
changeset
 | 
269  | 
((Binding.concealed Binding.empty, []), random_def)) random_defs')  | 
| 
31595
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
270  | 
|> snd  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59154 
diff
changeset
 | 
271  | 
|> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt [])  | 
| 
31595
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
272  | 
end;  | 
| 
 
bd2f7211a420
first running version of qc generators for datatypes
 
haftmann 
parents: 
31485 
diff
changeset
 | 
273  | 
|
| 59154 | 274  | 
|
| 31950 | 275  | 
(** building and compiling generator expressions **)  | 
276  | 
||
| 59154 | 277  | 
structure Data = Proof_Data  | 
| 
41472
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
40911 
diff
changeset
 | 
278  | 
(  | 
| 59154 | 279  | 
type T =  | 
280  | 
(unit -> Code_Numeral.natural -> bool -> Code_Numeral.natural -> seed ->  | 
|
281  | 
(bool * term list) option * seed) *  | 
|
282  | 
(unit -> Code_Numeral.natural -> bool -> Code_Numeral.natural -> seed ->  | 
|
283  | 
((bool * term list) option * (bool list * bool)) * seed);  | 
|
284  | 
val empty: T =  | 
|
285  | 
(fn () => raise Fail "counterexample",  | 
|
286  | 
fn () => raise Fail "counterexample_report");  | 
|
287  | 
fun init _ = empty;  | 
|
| 
39388
 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 
haftmann 
parents: 
39253 
diff
changeset
 | 
288  | 
);  | 
| 31950 | 289  | 
|
| 59154 | 290  | 
val get_counterexample = #1 o Data.get;  | 
291  | 
val get_counterexample_report = #2 o Data.get;  | 
|
292  | 
||
293  | 
val put_counterexample = Data.map o @{apply 2(1)} o K;
 | 
|
294  | 
val put_counterexample_report = Data.map o @{apply 2(2)} o K;
 | 
|
| 
35378
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35166 
diff
changeset
 | 
295  | 
|
| 31950 | 296  | 
val target = "Quickcheck";  | 
297  | 
||
| 50046 | 298  | 
fun mk_generator_expr ctxt (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
 | 
299  | 
let  | 
| 42361 | 300  | 
val thy = Proof_Context.theory_of ctxt  | 
| 44241 | 301  | 
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
 | 
302  | 
val Ts = (map snd o fst o strip_abs) prop  | 
| 31950 | 303  | 
val bound_max = length Ts - 1;  | 
304  | 
val bounds = map_index (fn (i, ty) =>  | 
|
305  | 
(2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) Ts;  | 
|
306  | 
val result = list_comb (prop, map (fn (i, _, _, _) => Bound i) bounds);  | 
|
| 67149 | 307  | 
val terms = HOLogic.mk_list \<^typ>\<open>term\<close> (map (fn (_, i, _, _) => Bound i $ \<^term>\<open>()\<close>) bounds);  | 
| 50046 | 308  | 
val ([genuine_only_name], _) = Variable.variant_fixes ["genuine_only"] ctxt  | 
| 67149 | 309  | 
val genuine_only = Free (genuine_only_name, \<^typ>\<open>bool\<close>)  | 
310  | 
val none_t = Const (\<^const_name>\<open>None\<close>, resultT)  | 
|
| 
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
 | 
311  | 
val check = Quickcheck_Common.mk_safe_if genuine_only none_t (result, none_t,  | 
| 67149 | 312  | 
fn genuine => \<^term>\<open>Some :: bool \<times> term list => (bool \<times> term list) option\<close> $  | 
| 
45721
 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 
bulwahn 
parents: 
45719 
diff
changeset
 | 
313  | 
HOLogic.mk_prod (Quickcheck_Common.reflect_bool genuine, terms))  | 
| 67149 | 314  | 
val return = HOLogic.pair_const resultT \<^typ>\<open>Random.seed\<close>;  | 
| 31950 | 315  | 
fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT);  | 
| 67149 | 316  | 
fun mk_termtyp T = HOLogic.mk_prodT (T, \<^typ>\<open>unit \<Rightarrow> term\<close>);  | 
317  | 
fun mk_scomp T1 T2 sT f g = Const (\<^const_name>\<open>scomp\<close>,  | 
|
| 31950 | 318  | 
liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g;  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61125 
diff
changeset
 | 
319  | 
fun mk_case_prod T = Sign.mk_const thy  | 
| 67149 | 320  | 
(\<^const_name>\<open>case_prod\<close>, [T, \<^typ>\<open>unit \<Rightarrow> term\<close>, liftT resultT \<^typ>\<open>Random.seed\<close>]);  | 
| 31950 | 321  | 
fun mk_scomp_split T t t' =  | 
| 67149 | 322  | 
mk_scomp (mk_termtyp T) resultT \<^typ>\<open>Random.seed\<close> t  | 
323  | 
        (mk_case_prod T $ Abs ("", T, Abs ("", \<^typ>\<open>unit => term\<close>, t')));
 | 
|
| 31950 | 324  | 
fun mk_bindclause (_, _, i, T) = mk_scomp_split T  | 
| 67149 | 325  | 
(Sign.mk_const thy (\<^const_name>\<open>Quickcheck_Random.random\<close>, [T]) $ Bound i);  | 
| 
45754
 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 
bulwahn 
parents: 
45728 
diff
changeset
 | 
326  | 
in  | 
| 
 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 
bulwahn 
parents: 
45728 
diff
changeset
 | 
327  | 
lambda genuine_only  | 
| 67149 | 328  | 
      (Abs ("n", \<^typ>\<open>natural\<close>, fold_rev mk_bindclause bounds (return $ check true)))
 | 
| 
45754
 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 
bulwahn 
parents: 
45728 
diff
changeset
 | 
329  | 
end;  | 
| 31950 | 330  | 
|
| 50046 | 331  | 
fun mk_reporting_generator_expr ctxt (t, _) =  | 
| 
35378
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35166 
diff
changeset
 | 
332  | 
let  | 
| 42361 | 333  | 
val thy = Proof_Context.theory_of ctxt  | 
| 67149 | 334  | 
val resultT = \<^typ>\<open>(bool \<times> term list) option \<times> (bool list \<times> bool)\<close>  | 
| 44241 | 335  | 
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
 | 
336  | 
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
 | 
337  | 
val bound_max = length Ts - 1  | 
| 
35378
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35166 
diff
changeset
 | 
338  | 
val bounds = map_index (fn (i, ty) =>  | 
| 
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35166 
diff
changeset
 | 
339  | 
(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
 | 
340  | 
val prop' = betapplys (prop, map (fn (i, _, _, _) => Bound i) bounds);  | 
| 67149 | 341  | 
val terms = HOLogic.mk_list \<^typ>\<open>term\<close> (map (fn (_, i, _, _) => Bound i $ \<^term>\<open>()\<close>) 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
 | 
342  | 
val (assms, concl) = Quickcheck_Common.strip_imp prop'  | 
| 67149 | 343  | 
val return = HOLogic.pair_const resultT \<^typ>\<open>Random.seed\<close>;  | 
| 
35378
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35166 
diff
changeset
 | 
344  | 
fun mk_assms_report i =  | 
| 67149 | 345  | 
HOLogic.mk_prod (\<^term>\<open>None :: (bool \<times> term list) option\<close>,  | 
| 
38553
 
56965d8e4e11
use HOLogic.boolT and @{typ bool} more pervasively
 
haftmann 
parents: 
38549 
diff
changeset
 | 
346  | 
HOLogic.mk_prod (HOLogic.mk_list HOLogic.boolT  | 
| 67149 | 347  | 
(replicate i \<^term>\<open>True\<close> @ replicate (length assms - i) \<^term>\<open>False\<close>),  | 
348  | 
\<^term>\<open>False\<close>))  | 
|
| 
35378
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35166 
diff
changeset
 | 
349  | 
fun mk_concl_report b =  | 
| 67149 | 350  | 
HOLogic.mk_prod (HOLogic.mk_list HOLogic.boolT (replicate (length assms) \<^term>\<open>True\<close>),  | 
| 
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
 | 
351  | 
Quickcheck_Common.reflect_bool b)  | 
| 50046 | 352  | 
val ([genuine_only_name], _) = Variable.variant_fixes ["genuine_only"] ctxt  | 
| 67149 | 353  | 
val genuine_only = Free (genuine_only_name, \<^typ>\<open>bool\<close>)  | 
354  | 
val none_t = HOLogic.mk_prod (\<^term>\<open>None :: (bool \<times> term list) option\<close>, mk_concl_report 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
 | 
355  | 
val concl_check = Quickcheck_Common.mk_safe_if genuine_only none_t (concl, none_t,  | 
| 67149 | 356  | 
fn genuine => HOLogic.mk_prod (\<^term>\<open>Some :: bool \<times> term list => (bool \<times> term list) option\<close> $  | 
| 
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
 | 
357  | 
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
 | 
358  | 
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
 | 
359  | 
(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
 | 
360  | 
(map_index I assms) concl_check  | 
| 
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35166 
diff
changeset
 | 
361  | 
fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT);  | 
| 67149 | 362  | 
fun mk_termtyp T = HOLogic.mk_prodT (T, \<^typ>\<open>unit \<Rightarrow> term\<close>);  | 
363  | 
fun mk_scomp T1 T2 sT f g = Const (\<^const_name>\<open>scomp\<close>,  | 
|
| 
35378
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35166 
diff
changeset
 | 
364  | 
liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g;  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61125 
diff
changeset
 | 
365  | 
fun mk_case_prod T = Sign.mk_const thy  | 
| 67149 | 366  | 
(\<^const_name>\<open>case_prod\<close>, [T, \<^typ>\<open>unit \<Rightarrow> term\<close>, liftT resultT \<^typ>\<open>Random.seed\<close>]);  | 
| 
35378
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35166 
diff
changeset
 | 
367  | 
fun mk_scomp_split T t t' =  | 
| 67149 | 368  | 
mk_scomp (mk_termtyp T) resultT \<^typ>\<open>Random.seed\<close> t  | 
369  | 
        (mk_case_prod T $ Abs ("", T, Abs ("", \<^typ>\<open>unit \<Rightarrow> term\<close>, t')));
 | 
|
| 
35378
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35166 
diff
changeset
 | 
370  | 
fun mk_bindclause (_, _, i, T) = mk_scomp_split T  | 
| 67149 | 371  | 
(Sign.mk_const thy (\<^const_name>\<open>Quickcheck_Random.random\<close>, [T]) $ Bound i);  | 
| 
35378
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35166 
diff
changeset
 | 
372  | 
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
 | 
373  | 
lambda genuine_only  | 
| 67149 | 374  | 
      (Abs ("n", \<^typ>\<open>natural\<close>, fold_rev mk_bindclause bounds (return $ check true)))
 | 
| 
35378
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35166 
diff
changeset
 | 
375  | 
end  | 
| 
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35166 
diff
changeset
 | 
376  | 
|
| 
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
 | 
377  | 
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
 | 
378  | 
((mk_generator_expr,  | 
| 67149 | 379  | 
absdummy \<^typ>\<open>bool\<close> (absdummy \<^typ>\<open>natural\<close>  | 
380  | 
\<^term>\<open>Pair None :: Random.seed \<Rightarrow> (bool \<times> term list) option \<times> Random.seed\<close>)),  | 
|
381  | 
\<^typ>\<open>bool \<Rightarrow> natural \<Rightarrow> Random.seed \<Rightarrow> (bool \<times> term list) option \<times> Random.seed\<close>)  | 
|
| 
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
 | 
382  | 
|
| 
 
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  | 
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
 | 
384  | 
((mk_reporting_generator_expr,  | 
| 67149 | 385  | 
absdummy \<^typ>\<open>bool\<close> (absdummy \<^typ>\<open>natural\<close>  | 
386  | 
\<^term>\<open>Pair (None, ([], False)) :: Random.seed \<Rightarrow>  | 
|
387  | 
((bool \<times> term list) option \<times> (bool list \<times> bool)) \<times> Random.seed\<close>)),  | 
|
388  | 
\<^typ>\<open>bool \<Rightarrow> natural \<Rightarrow> Random.seed \<Rightarrow> ((bool \<times> term list) option \<times> (bool list \<times> bool)) \<times> Random.seed\<close>)  | 
|
| 
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
 | 
389  | 
|
| 
 
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
 | 
390  | 
|
| 
40911
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
391  | 
(* single quickcheck report *)  | 
| 
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
392  | 
|
| 
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
393  | 
datatype single_report = Run of bool list * bool | MatchExc  | 
| 
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
394  | 
|
| 
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
395  | 
fun collect_single_report single_report  | 
| 
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
396  | 
    (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
 | 
397  | 
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
 | 
398  | 
case single_report  | 
| 
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
399  | 
of MatchExc =>  | 
| 
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
400  | 
    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
 | 
401  | 
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
 | 
402  | 
| Run (assms, concl) =>  | 
| 
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
403  | 
    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
 | 
404  | 
satisfied_assms =  | 
| 
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
405  | 
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
 | 
406  | 
(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
 | 
407  | 
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
 | 
408  | 
|
| 
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
409  | 
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
 | 
410  | 
satisfied_assms = [], positive_concl_tests = 0 }  | 
| 
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
411  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
412  | 
fun compile_generator_expr_raw ctxt ts =  | 
| 31950 | 413  | 
let  | 
| 
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;  | 
| 59154 | 419  | 
val compile =  | 
420  | 
Code_Runtime.dynamic_value_strict  | 
|
421  | 
(get_counterexample_report, put_counterexample_report,  | 
|
422  | 
"Random_Generators.put_counterexample_report")  | 
|
423  | 
ctxt (SOME target)  | 
|
424  | 
(fn proc => fn g => fn c => fn b => fn s =>  | 
|
425  | 
g c b s #>> (apfst o Option.map o apsnd o map) proc)  | 
|
426  | 
t' [];  | 
|
| 
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
 | 
427  | 
fun single_tester c b s = compile c b s |> Random_Engine.run  | 
| 50046 | 428  | 
fun iterate_and_collect _ _ 0 report = (NONE, report)  | 
| 
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
 | 
429  | 
| 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
 | 
430  | 
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
 | 
431  | 
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
 | 
432  | 
val report = collect_single_report single_report report  | 
| 
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
433  | 
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
 | 
434  | 
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
 | 
435  | 
| SOME q => (SOME q, report)  | 
| 
40911
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
436  | 
end  | 
| 
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
437  | 
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
 | 
438  | 
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
 | 
439  | 
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
 | 
440  | 
end  | 
| 
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
441  | 
else  | 
| 
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
442  | 
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
 | 
443  | 
val t' = mk_parametric_generator_expr ctxt ts;  | 
| 59154 | 444  | 
val compile =  | 
445  | 
Code_Runtime.dynamic_value_strict  | 
|
446  | 
(get_counterexample, put_counterexample, "Random_Generators.put_counterexample")  | 
|
447  | 
ctxt (SOME target)  | 
|
448  | 
(fn proc => fn g => fn c => fn b => fn s =>  | 
|
449  | 
g c b s #>> (Option.map o apsnd o map) proc)  | 
|
450  | 
t' [];  | 
|
| 
45754
 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 
bulwahn 
parents: 
45728 
diff
changeset
 | 
451  | 
fun single_tester c b s = compile c b s |> Random_Engine.run  | 
| 50046 | 452  | 
fun iterate _ _ 0 = NONE  | 
| 
45754
 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 
bulwahn 
parents: 
45728 
diff
changeset
 | 
453  | 
| iterate genuine_only (card, size) j =  | 
| 
 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 
bulwahn 
parents: 
45728 
diff
changeset
 | 
454  | 
case single_tester card genuine_only size of  | 
| 
 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 
bulwahn 
parents: 
45728 
diff
changeset
 | 
455  | 
NONE => iterate genuine_only (card, size) (j - 1)  | 
| 
 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 
bulwahn 
parents: 
45728 
diff
changeset
 | 
456  | 
| SOME q => SOME q  | 
| 
40911
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
457  | 
in  | 
| 
45754
 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 
bulwahn 
parents: 
45728 
diff
changeset
 | 
458  | 
fn genuine_only => fn [card, size] =>  | 
| 
 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 
bulwahn 
parents: 
45728 
diff
changeset
 | 
459  | 
(rpair NONE (iterate genuine_only (card, size) iterations))  | 
| 
40911
 
7febf76e0a69
moving iteration of tests to the testers in quickcheck
 
bulwahn 
parents: 
40644 
diff
changeset
 | 
460  | 
end  | 
| 
35378
 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 
bulwahn 
parents: 
35166 
diff
changeset
 | 
461  | 
end;  | 
| 31950 | 462  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
463  | 
fun compile_generator_expr ctxt ts =  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
464  | 
let  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
465  | 
val compiled = compile_generator_expr_raw ctxt ts  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
466  | 
in fn genuine_only => fn [card, size] =>  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
467  | 
compiled genuine_only [Code_Numeral.natural_of_integer card, Code_Numeral.natural_of_integer size]  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
468  | 
end;  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
51126 
diff
changeset
 | 
469  | 
|
| 67149 | 470  | 
val size_types = [\<^type_name>\<open>Enum.finite_1\<close>, \<^type_name>\<open>Enum.finite_2\<close>,  | 
471  | 
\<^type_name>\<open>Enum.finite_3\<close>, \<^type_name>\<open>Enum.finite_4\<close>, \<^type_name>\<open>Enum.finite_5\<close>];  | 
|
| 50818 | 472  | 
|
473  | 
fun size_matters_for _ Ts =  | 
|
474  | 
not (forall (fn Type (tyco, []) => member (op =) size_types tyco | _ => false) Ts);  | 
|
| 
46331
 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 
bulwahn 
parents: 
45940 
diff
changeset
 | 
475  | 
|
| 
 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 
bulwahn 
parents: 
45940 
diff
changeset
 | 
476  | 
val test_goals =  | 
| 
 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 
bulwahn 
parents: 
45940 
diff
changeset
 | 
477  | 
  Quickcheck_Common.generator_test_goal_terms ("random", (size_matters_for, compile_generator_expr));
 | 
| 62979 | 478  | 
|
| 
43875
 
485d2ad43528
adding random, exhaustive and SML quickcheck as testers
 
bulwahn 
parents: 
43333 
diff
changeset
 | 
479  | 
|
| 31260 | 480  | 
(** setup **)  | 
481  | 
||
| 67149 | 482  | 
val active = Attrib.setup_config_bool \<^binding>\<open>quickcheck_random_active\<close> (K false);  | 
| 
43878
 
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
 
bulwahn 
parents: 
43877 
diff
changeset
 | 
483  | 
|
| 58826 | 484  | 
val _ =  | 
485  | 
Theory.setup  | 
|
| 67149 | 486  | 
(Quickcheck_Common.datatype_interpretation \<^plugin>\<open>quickcheck_random\<close>  | 
487  | 
(\<^sort>\<open>random\<close>, instantiate_random_datatype) #>  | 
|
| 58826 | 488  | 
    Context.theory_map (Quickcheck.add_tester ("random", (active, test_goals))));
 | 
| 31260 | 489  | 
|
490  | 
end;  |