| author | wenzelm | 
| Sun, 14 Nov 2010 14:05:08 +0100 | |
| changeset 40534 | 9e196062bf88 | 
| parent 40420 | 552563ea3304 | 
| child 40639 | f1f0e6adca0a | 
| permissions | -rw-r--r-- | 
| 
40420
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Tools/smallvalue_generators.ML  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
2  | 
Author: Lukas Bulwahn, TU Muenchen  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
3  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
4  | 
Generators for small values for various types.  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
7  | 
signature SMALLVALUE_GENERATORS =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
8  | 
sig  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
9  | 
val ensure_smallvalue_datatype: Datatype.config -> string list -> theory -> theory  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
10  | 
val compile_generator_expr:  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
11  | 
Proof.context -> term -> int -> term list option * (bool list * bool)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
12  | 
val put_counterexample: (unit -> int -> term list option)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
13  | 
-> Proof.context -> Proof.context  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
14  | 
val setup: theory -> theory  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
15  | 
end;  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
16  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
17  | 
structure Smallvalue_Generators : SMALLVALUE_GENERATORS =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
18  | 
struct  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
19  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
20  | 
(** general term functions **)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
21  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
22  | 
fun dest_funT (Type ("fun",[S, T])) = (S, T)
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
23  | 
  | dest_funT T = raise TYPE ("dest_funT", [T], [])
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
24  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
25  | 
fun mk_fun_comp (t, u) =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
26  | 
let  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
27  | 
val (_, B) = dest_funT (fastype_of t)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
28  | 
val (C, A) = dest_funT (fastype_of u)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
29  | 
in  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
30  | 
    Const(@{const_name "Fun.comp"}, (A --> B) --> (C --> A) --> C --> B) $ t $ u
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
31  | 
end;  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
32  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
33  | 
fun mk_measure f =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
34  | 
let  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
35  | 
    val Type ("fun", [T, @{typ nat}]) = fastype_of f 
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
36  | 
in  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
37  | 
    Const (@{const_name Wellfounded.measure},
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
38  | 
      (T --> @{typ nat}) --> HOLogic.mk_prodT (T, T) --> @{typ bool})
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
39  | 
$ f  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
40  | 
end  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
41  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
42  | 
fun mk_sumcases rT f (Type (@{type_name Sum_Type.sum}, [TL, TR])) =
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
43  | 
let  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
44  | 
val lt = mk_sumcases rT f TL  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
45  | 
val rt = mk_sumcases rT f TR  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
46  | 
in  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
47  | 
SumTree.mk_sumcase TL TR rT lt rt  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
48  | 
end  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
49  | 
| mk_sumcases _ f T = f T  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
50  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
52  | 
(** abstract syntax **)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
53  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
54  | 
val size = @{term "i :: code_numeral"}
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
55  | 
val size_pred = @{term "(i :: code_numeral) - 1"}
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
56  | 
val size_ge_zero = @{term "(i :: code_numeral) > 0"}
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
57  | 
fun test_function T = Free ("f", T --> @{typ "term list option"})
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
58  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
59  | 
fun mk_none_continuation (x, y) =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
60  | 
let  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
61  | 
    val (T as Type(@{type_name "option"}, [T'])) = fastype_of x
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
62  | 
in  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
63  | 
    Const (@{const_name Option.option_case}, T --> (T' --> T) --> T --> T)
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
64  | 
      $ y $ Const (@{const_name Some}, T' --> T) $ x
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
65  | 
end  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
66  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
67  | 
(** datatypes **)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
68  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
69  | 
(* constructing smallvalue generator instances on datatypes *)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
70  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
71  | 
exception FUNCTION_TYPE;  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
72  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
73  | 
val smallN = "small";  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
74  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
75  | 
fun smallT T = (T --> @{typ "Code_Evaluation.term list option"}) --> @{typ code_numeral}
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
76  | 
  --> @{typ "Code_Evaluation.term list option"}
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
77  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
78  | 
fun mk_equations thy descr vs tycos (names, auxnames) (Ts, Us) =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
79  | 
let  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
80  | 
val smallsN = map (prefix (smallN ^ "_")) (names @ auxnames);  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
81  | 
val smalls = map2 (fn name => fn T => Free (name, smallT T))  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
82  | 
smallsN (Ts @ Us)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
83  | 
fun mk_small_call T =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
84  | 
let  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
85  | 
        val small = Const (@{const_name "Smallcheck.small_class.small"}, smallT T)        
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
86  | 
in  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
87  | 
(T, (fn t => small $ absdummy (T, t) $ size_pred))  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
88  | 
end  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
89  | 
fun mk_small_aux_call fTs (k, _) (tyco, Ts) =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
90  | 
let  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
91  | 
val T = Type (tyco, Ts)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
92  | 
val _ = if not (null fTs) then raise FUNCTION_TYPE else ()  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
93  | 
val small = nth smalls k  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
94  | 
in  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
95  | 
(T, (fn t => small $ absdummy (T, t) $ size_pred))  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
96  | 
end  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
97  | 
fun mk_consexpr simpleT (c, xs) =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
98  | 
let  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
99  | 
val (Ts, fns) = split_list xs  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
100  | 
val constr = Const (c, Ts ---> simpleT)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
101  | 
val bounds = map Bound (((length xs) - 1) downto 0)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
102  | 
val start_term = test_function simpleT $ (list_comb (constr, bounds))  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
103  | 
in fold_rev (fn f => fn t => f t) fns start_term end  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
104  | 
fun mk_rhs exprs =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
105  | 
        @{term "If :: bool => term list option => term list option => term list option"}
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
106  | 
            $ size_ge_zero $ (foldr1 mk_none_continuation exprs) $ @{term "None :: term list option"}
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
107  | 
val rhss =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
108  | 
Datatype_Aux.interpret_construction descr vs  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
109  | 
        { atyp = mk_small_call, dtyp = mk_small_aux_call }
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
110  | 
|> (map o apfst) Type  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
111  | 
|> map (fn (T, cs) => map (mk_consexpr T) cs)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
112  | 
|> map mk_rhs  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
113  | 
val lhss = map2 (fn t => fn T => t $ test_function T $ size) smalls (Ts @ Us);  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
114  | 
val eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (lhss ~~ rhss)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
115  | 
in  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
116  | 
(Ts @ Us ~~ (smallsN ~~ eqs))  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
117  | 
end  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
118  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
119  | 
val less_int_pred = @{lemma "i > 0 ==> Code_Numeral.nat_of ((i :: code_numeral) - 1) < Code_Numeral.nat_of i" by auto}
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
120  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
121  | 
fun gen_inst_state_tac ctxt rel st =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
122  | 
case Term.add_vars (prop_of st) [] of  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
123  | 
[v as (_, T)] =>  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
124  | 
let  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
125  | 
val cert = Thm.cterm_of (ProofContext.theory_of ctxt)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
126  | 
val rel' = cert rel  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
127  | 
val st' = Thm.incr_indexes (#maxidx (Thm.rep_cterm rel') + 1) st (*FIXME??*)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
128  | 
in  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
129  | 
PRIMITIVE (Drule.cterm_instantiate [(cert (Var v), rel')]) st'  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
130  | 
end  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
131  | 
| _ => Seq.empty;  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
132  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
133  | 
fun instantiate_smallvalue_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
134  | 
let  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
135  | 
val _ = Datatype_Aux.message config "Creating smallvalue generators ...";  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
136  | 
val eqs = mk_equations thy descr vs tycos (names, auxnames) (Ts, Us)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
137  | 
fun my_relation_tac ctxt st =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
138  | 
let  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
139  | 
val ((_ $ (_ $ rel)) :: tl) = prems_of st  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
140  | 
val domT = (HOLogic.dest_setT (fastype_of rel))  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
141  | 
        fun mk_single_measure T = mk_fun_comp (@{term "Code_Numeral.nat_of"},
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
142  | 
            Const (@{const_name "Product_Type.snd"}, T --> @{typ "code_numeral"}))
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
143  | 
        val measure = mk_measure (mk_sumcases @{typ nat} mk_single_measure domT)
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
144  | 
in  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
145  | 
(Function_Common.apply_termination_rule ctxt 1  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
146  | 
THEN gen_inst_state_tac ctxt measure) st  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
147  | 
end  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
148  | 
fun termination_tac ctxt =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
149  | 
my_relation_tac ctxt  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
150  | 
      THEN rtac @{thm wf_measure} 1
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
151  | 
THEN (REPEAT_DETERM (Simplifier.asm_full_simp_tac  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
152  | 
        (HOL_basic_ss addsimps [@{thm in_measure}, @{thm o_def}, @{thm snd_conv},
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
153  | 
         @{thm nat_mono_iff}, less_int_pred] @ @{thms sum.cases}) 1))
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
154  | 
fun pat_completeness_auto ctxt =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
155  | 
Pat_Completeness.pat_completeness_tac ctxt 1  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
156  | 
THEN auto_tac (clasimpset_of ctxt)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
157  | 
in  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
158  | 
thy  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
159  | 
    |> Class.instantiation (tycos, vs, @{sort small})
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
160  | 
|> Function.add_function  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
161  | 
(map (fn (T, (name, _)) =>  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
162  | 
Syntax.no_syn (Binding.conceal (Binding.name name), SOME (smallT T))) eqs)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
163  | 
(map (pair (apfst Binding.conceal Attrib.empty_binding) o snd o snd) eqs)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
164  | 
Function_Common.default_config pat_completeness_auto  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
165  | 
|> snd  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
166  | 
|> Local_Theory.restore  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
167  | 
|> (fn lthy => Function.prove_termination NONE (termination_tac lthy) lthy)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
168  | 
|> snd  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
169  | 
|> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
170  | 
end;  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
171  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
172  | 
fun ensure_smallvalue_datatype config raw_tycos thy =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
173  | 
let  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
174  | 
val algebra = Sign.classes_of thy;  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
175  | 
val (descr, raw_vs, tycos, prfx, (names, auxnames), raw_TUs) =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
176  | 
Datatype.the_descr thy raw_tycos;  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
177  | 
val typerep_vs = (map o apsnd)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
178  | 
      (curry (Sorts.inter_sort algebra) @{sort typerep}) raw_vs;
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
179  | 
    val smallvalue_insts = (map (rpair @{sort small}) o flat o maps snd o maps snd)
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
180  | 
(Datatype_Aux.interpret_construction descr typerep_vs  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
181  | 
        { atyp = single, dtyp = (K o K o K) [] });
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
182  | 
    (*val term_of_insts = (map (rpair @{sort term_of}) o flat o maps snd o maps snd)
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
183  | 
(Datatype_Aux.interpret_construction descr typerep_vs  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
184  | 
        { atyp = K [], dtyp = K o K });*)
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
185  | 
val has_inst = exists (fn tyco =>  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
186  | 
      can (Sorts.mg_domain algebra tyco) @{sort small}) tycos;
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
187  | 
in if has_inst then thy  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
188  | 
else case Quickcheck_Generators.perhaps_constrain thy smallvalue_insts typerep_vs  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
189  | 
of SOME constrain => (instantiate_smallvalue_datatype config descr  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
190  | 
(map constrain typerep_vs) tycos prfx (names, auxnames)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
191  | 
((pairself o map o map_atyps) (fn TFree v => TFree (constrain v)) raw_TUs) thy  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
192  | 
handle FUNCTION_TYPE =>  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
193  | 
(Datatype_Aux.message config  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
194  | 
"Creation of smallvalue generators failed because the datatype contains a function type";  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
195  | 
thy))  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
196  | 
| NONE => thy  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
197  | 
end;  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
198  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
199  | 
(** building and compiling generator expressions **)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
200  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
201  | 
structure Counterexample = Proof_Data (  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
202  | 
type T = unit -> int -> term list option  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
203  | 
fun init _ () = error "Counterexample"  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
204  | 
);  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
205  | 
val put_counterexample = Counterexample.put;  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
206  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
207  | 
val target = "Quickcheck";  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
208  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
209  | 
fun mk_generator_expr thy prop Ts =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
210  | 
let  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
211  | 
val bound_max = length Ts - 1;  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
212  | 
val bounds = map Bound (bound_max downto 0)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
213  | 
val result = list_comb (prop, bounds);  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
214  | 
    val terms = HOLogic.mk_list @{typ term} (map2 HOLogic.mk_term_of Ts bounds);
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
215  | 
val check =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
216  | 
      @{term "Smallcheck.catch_match :: term list option => term list option => term list option"} $
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
217  | 
        (@{term "If :: bool => term list option => term list option => term list option"}
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
218  | 
        $ result $ @{term "None :: term list option"}
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
219  | 
        $ (@{term "Some :: term list => term list option"} $ terms))
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
220  | 
      $ @{term "None :: term list option"};
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
221  | 
fun mk_small_closure (depth, T) t =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
222  | 
      Const (@{const_name "Smallcheck.small_class.small"}, smallT T)
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
223  | 
$ absdummy (T, t) $ depth  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
224  | 
  in Abs ("d", @{typ code_numeral}, fold_rev mk_small_closure (rev bounds ~~ Ts) check) end
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
225  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
226  | 
fun compile_generator_expr ctxt t =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
227  | 
let  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
228  | 
val Ts = (map snd o fst o strip_abs) t;  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
229  | 
val thy = ProofContext.theory_of ctxt  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
230  | 
in if Quickcheck.report ctxt then  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
231  | 
error "Compilation with reporting facility is not supported"  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
232  | 
else  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
233  | 
let  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
234  | 
val t' = mk_generator_expr thy t Ts;  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
235  | 
val compile = Code_Runtime.dynamic_value_strict  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
236  | 
(Counterexample.get, put_counterexample, "Smallvalue_Generators.put_counterexample")  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
237  | 
thy (SOME target) (fn proc => fn g => g #> (Option.map o map) proc) t' [];  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
238  | 
val dummy_report = ([], false)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
239  | 
in compile #> rpair dummy_report end  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
240  | 
end;  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
241  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
242  | 
(** setup **)  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
243  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
244  | 
val setup =  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
245  | 
Datatype.interpretation ensure_smallvalue_datatype  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
246  | 
#> Context.theory_map  | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
247  | 
    (Quickcheck.add_generator ("small", compile_generator_expr));
 | 
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
248  | 
|
| 
 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 
bulwahn 
parents:  
diff
changeset
 | 
249  | 
end;  |