author | haftmann |
Fri, 17 Apr 2009 14:29:54 +0200 | |
changeset 30945 | 0418e9bffbba |
parent 30364 | 577edc39b501 |
child 31135 | e2d777dcf161 |
permissions | -rw-r--r-- |
29132 | 1 |
(* Author: Florian Haftmann, TU Muenchen *) |
26265 | 2 |
|
29808
b8b9d529663b
split of already properly working part of Quickcheck infrastructure
haftmann
parents:
29579
diff
changeset
|
3 |
header {* Experimental counterexample generators *} |
26265 | 4 |
|
29808
b8b9d529663b
split of already properly working part of Quickcheck infrastructure
haftmann
parents:
29579
diff
changeset
|
5 |
theory Quickcheck_Generators |
b8b9d529663b
split of already properly working part of Quickcheck infrastructure
haftmann
parents:
29579
diff
changeset
|
6 |
imports Quickcheck State_Monad |
26265 | 7 |
begin |
8 |
||
29808
b8b9d529663b
split of already properly working part of Quickcheck infrastructure
haftmann
parents:
29579
diff
changeset
|
9 |
subsection {* Datatypes *} |
29132 | 10 |
|
11 |
definition |
|
12 |
collapse :: "('a \<Rightarrow> ('a \<Rightarrow> 'b \<times> 'a) \<times> 'a) \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where |
|
13 |
"collapse f = (do g \<leftarrow> f; g done)" |
|
14 |
||
15 |
ML {* |
|
16 |
structure StateMonad = |
|
17 |
struct |
|
18 |
||
19 |
fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT); |
|
20 |
fun liftT' sT = sT --> sT; |
|
21 |
||
22 |
fun return T sT x = Const (@{const_name return}, T --> liftT T sT) $ x; |
|
23 |
||
24 |
fun scomp T1 T2 sT f g = Const (@{const_name scomp}, |
|
25 |
liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g; |
|
26 |
||
27 |
end; |
|
28 |
*} |
|
29 |
||
30 |
lemma random'_if: |
|
31 |
fixes random' :: "index \<Rightarrow> index \<Rightarrow> seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> seed" |
|
32 |
assumes "random' 0 j = (\<lambda>s. undefined)" |
|
33 |
and "\<And>i. random' (Suc_index i) j = rhs2 i" |
|
34 |
shows "random' i j s = (if i = 0 then undefined else rhs2 (i - 1) s)" |
|
35 |
by (cases i rule: index.exhaust) (insert assms, simp_all) |
|
36 |
||
37 |
setup {* |
|
38 |
let |
|
39 |
exception REC of string; |
|
40 |
exception TYP of string; |
|
41 |
fun mk_collapse thy ty = Sign.mk_const thy |
|
42 |
(@{const_name collapse}, [@{typ seed}, ty]); |
|
43 |
fun term_ty ty = HOLogic.mk_prodT (ty, @{typ "unit \<Rightarrow> term"}); |
|
44 |
fun mk_split thy ty ty' = Sign.mk_const thy |
|
45 |
(@{const_name split}, [ty, @{typ "unit \<Rightarrow> term"}, StateMonad.liftT (term_ty ty') @{typ seed}]); |
|
46 |
fun mk_scomp_split thy ty ty' t t' = |
|
47 |
StateMonad.scomp (term_ty ty) (term_ty ty') @{typ seed} t |
|
48 |
(mk_split thy ty ty' $ Abs ("", ty, Abs ("", @{typ "unit \<Rightarrow> term"}, t'))) |
|
49 |
fun mk_cons thy this_ty (c, args) = |
|
50 |
let |
|
51 |
val tys = map (fst o fst) args; |
|
52 |
val c_ty = tys ---> this_ty; |
|
53 |
val c = Const (c, tys ---> this_ty); |
|
54 |
val t_indices = map (curry ( op * ) 2) (length tys - 1 downto 0); |
|
55 |
val c_indices = map (curry ( op + ) 1) t_indices; |
|
56 |
val c_t = list_comb (c, map Bound c_indices); |
|
57 |
val t_t = Abs ("", @{typ unit}, Eval.mk_term Free Typerep.typerep |
|
58 |
(list_comb (c, map (fn k => Bound (k + 1)) t_indices)) |
|
59 |
|> map_aterms (fn t as Bound _ => t $ @{term "()"} | t => t)); |
|
60 |
val return = StateMonad.return (term_ty this_ty) @{typ seed} |
|
61 |
(HOLogic.mk_prod (c_t, t_t)); |
|
62 |
val t = fold_rev (fn ((ty, _), random) => |
|
63 |
mk_scomp_split thy ty this_ty random) |
|
64 |
args return; |
|
65 |
val is_rec = exists (snd o fst) args; |
|
66 |
in (is_rec, t) end; |
|
67 |
fun mk_conss thy ty [] = NONE |
|
68 |
| mk_conss thy ty [(_, t)] = SOME t |
|
69 |
| mk_conss thy ty ts = SOME (mk_collapse thy (term_ty ty) $ |
|
70 |
(Sign.mk_const thy (@{const_name select}, [StateMonad.liftT (term_ty ty) @{typ seed}]) $ |
|
71 |
HOLogic.mk_list (StateMonad.liftT (term_ty ty) @{typ seed}) (map snd ts))); |
|
72 |
fun mk_clauses thy ty (tyco, (ts_rec, ts_atom)) = |
|
73 |
let |
|
74 |
val SOME t_atom = mk_conss thy ty ts_atom; |
|
75 |
in case mk_conss thy ty ts_rec |
|
76 |
of SOME t_rec => mk_collapse thy (term_ty ty) $ |
|
77 |
(Sign.mk_const thy (@{const_name select_default}, [StateMonad.liftT (term_ty ty) @{typ seed}]) $ |
|
78 |
@{term "i\<Colon>index"} $ t_rec $ t_atom) |
|
79 |
| NONE => t_atom |
|
80 |
end; |
|
81 |
fun mk_random_eqs thy vs tycos = |
|
82 |
let |
|
83 |
val this_ty = Type (hd tycos, map TFree vs); |
|
84 |
val this_ty' = StateMonad.liftT (term_ty this_ty) @{typ seed}; |
|
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
85 |
val random_name = Long_Name.base_name @{const_name random}; |
29132 | 86 |
val random'_name = random_name ^ "_" ^ Class.type_name (hd tycos) ^ "'"; |
87 |
fun random ty = Sign.mk_const thy (@{const_name random}, [ty]); |
|
88 |
val random' = Free (random'_name, |
|
89 |
@{typ index} --> @{typ index} --> this_ty'); |
|
90 |
fun atom ty = if Sign.of_sort thy (ty, @{sort random}) |
|
91 |
then ((ty, false), random ty $ @{term "j\<Colon>index"}) |
|
92 |
else raise TYP |
|
93 |
("Will not generate random elements for type(s) " ^ quote (hd tycos)); |
|
94 |
fun dtyp tyco = ((this_ty, true), random' $ @{term "i\<Colon>index"} $ @{term "j\<Colon>index"}); |
|
95 |
fun rtyp tyco tys = raise REC |
|
96 |
("Will not generate random elements for mutual recursive type " ^ quote (hd tycos)); |
|
97 |
val rhss = DatatypePackage.construction_interpretation thy |
|
98 |
{ atom = atom, dtyp = dtyp, rtyp = rtyp } vs tycos |
|
99 |
|> (map o apsnd o map) (mk_cons thy this_ty) |
|
100 |
|> (map o apsnd) (List.partition fst) |
|
101 |
|> map (mk_clauses thy this_ty) |
|
102 |
val eqss = map ((apsnd o map) (HOLogic.mk_Trueprop o HOLogic.mk_eq) o (fn rhs => ((this_ty, random'), [ |
|
103 |
(random' $ @{term "0\<Colon>index"} $ @{term "j\<Colon>index"}, Abs ("s", @{typ seed}, |
|
104 |
Const (@{const_name undefined}, HOLogic.mk_prodT (term_ty this_ty, @{typ seed})))), |
|
105 |
(random' $ @{term "Suc_index i"} $ @{term "j\<Colon>index"}, rhs) |
|
106 |
]))) rhss; |
|
107 |
in eqss end; |
|
108 |
fun random_inst [tyco] thy = |
|
109 |
let |
|
110 |
val (raw_vs, _) = DatatypePackage.the_datatype_spec thy tyco; |
|
111 |
val vs = (map o apsnd) |
|
112 |
(curry (Sorts.inter_sort (Sign.classes_of thy)) @{sort random}) raw_vs; |
|
113 |
val ((this_ty, random'), eqs') = singleton (mk_random_eqs thy vs) tyco; |
|
114 |
val eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq) |
|
115 |
(Sign.mk_const thy (@{const_name random}, [this_ty]) $ @{term "i\<Colon>index"}, |
|
116 |
random' $ @{term "i\<Colon>index"} $ @{term "i\<Colon>index"}) |
|
117 |
val del_func = Attrib.internal (fn _ => Thm.declaration_attribute |
|
118 |
(fn thm => Context.mapping (Code.del_eqn thm) I)); |
|
119 |
fun add_code simps lthy = |
|
120 |
let |
|
121 |
val thy = ProofContext.theory_of lthy; |
|
122 |
val thm = @{thm random'_if} |
|
123 |
|> Drule.instantiate' [SOME (Thm.ctyp_of thy this_ty)] [SOME (Thm.cterm_of thy random')] |
|
124 |
|> (fn thm => thm OF simps) |
|
125 |
|> singleton (ProofContext.export lthy (ProofContext.init thy)); |
|
126 |
val c = (fst o dest_Const o fst o strip_comb o fst |
|
127 |
o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm; |
|
128 |
in |
|
129 |
lthy |
|
130 |
|> LocalTheory.theory (Code.del_eqns c |
|
29579 | 131 |
#> PureThy.add_thm ((Binding.name (fst (dest_Free random') ^ "_code"), thm), [Thm.kind_internal]) |
29132 | 132 |
#-> Code.add_eqn) |
133 |
end; |
|
134 |
in |
|
135 |
thy |
|
136 |
|> TheoryTarget.instantiation ([tyco], vs, @{sort random}) |
|
137 |
|> PrimrecPackage.add_primrec |
|
138 |
[(Binding.name (fst (dest_Free random')), SOME (snd (dest_Free random')), NoSyn)] |
|
139 |
(map (fn eq => ((Binding.empty, [del_func]), eq)) eqs') |
|
140 |
|-> add_code |
|
141 |
|> `(fn lthy => Syntax.check_term lthy eq) |
|
142 |
|-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq))) |
|
143 |
|> snd |
|
144 |
|> Class.prove_instantiation_instance (K (Class.intro_classes_tac [])) |
|
145 |
|> LocalTheory.exit_global |
|
146 |
end |
|
147 |
| random_inst tycos thy = raise REC |
|
148 |
("Will not generate random elements for mutual recursive type(s) " ^ commas (map quote tycos)); |
|
149 |
fun add_random_inst tycos thy = random_inst tycos thy |
|
150 |
handle REC msg => (warning msg; thy) |
|
151 |
| TYP msg => (warning msg; thy) |
|
152 |
in DatatypePackage.interpretation add_random_inst end |
|
153 |
*} |
|
154 |
||
29808
b8b9d529663b
split of already properly working part of Quickcheck infrastructure
haftmann
parents:
29579
diff
changeset
|
155 |
|
b8b9d529663b
split of already properly working part of Quickcheck infrastructure
haftmann
parents:
29579
diff
changeset
|
156 |
subsection {* Type @{typ int} *} |
29132 | 157 |
|
158 |
instantiation int :: random |
|
159 |
begin |
|
160 |
||
161 |
definition |
|
162 |
"random n = (do |
|
163 |
(b, _) \<leftarrow> random n; |
|
164 |
(m, t) \<leftarrow> random n; |
|
165 |
return (if b then (int m, \<lambda>u. Code_Eval.App (Code_Eval.Const (STR ''Int.int'') TYPEREP(nat \<Rightarrow> int)) (t ())) |
|
166 |
else (- int m, \<lambda>u. Code_Eval.App (Code_Eval.Const (STR ''HOL.uminus_class.uminus'') TYPEREP(int \<Rightarrow> int)) |
|
167 |
(Code_Eval.App (Code_Eval.Const (STR ''Int.int'') TYPEREP(nat \<Rightarrow> int)) (t ())))) |
|
168 |
done)" |
|
169 |
||
170 |
instance .. |
|
171 |
||
172 |
end |
|
173 |
||
26265 | 174 |
|
26267
ba710daf77a7
added combinator for interpretation of construction of datatype
haftmann
parents:
26265
diff
changeset
|
175 |
subsection {* Examples *} |
ba710daf77a7
added combinator for interpretation of construction of datatype
haftmann
parents:
26265
diff
changeset
|
176 |
|
28315 | 177 |
theorem "map g (map f xs) = map (g o f) xs" |
178 |
quickcheck [generator = code] |
|
179 |
by (induct xs) simp_all |
|
26325 | 180 |
|
28315 | 181 |
theorem "map g (map f xs) = map (f o g) xs" |
182 |
quickcheck [generator = code] |
|
183 |
oops |
|
184 |
||
185 |
theorem "rev (xs @ ys) = rev ys @ rev xs" |
|
186 |
quickcheck [generator = code] |
|
187 |
by simp |
|
26265 | 188 |
|
28315 | 189 |
theorem "rev (xs @ ys) = rev xs @ rev ys" |
190 |
quickcheck [generator = code] |
|
191 |
oops |
|
192 |
||
193 |
theorem "rev (rev xs) = xs" |
|
194 |
quickcheck [generator = code] |
|
195 |
by simp |
|
196 |
||
197 |
theorem "rev xs = xs" |
|
198 |
quickcheck [generator = code] |
|
199 |
oops |
|
200 |
||
201 |
primrec app :: "('a \<Rightarrow> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a" where |
|
202 |
"app [] x = x" |
|
203 |
| "app (f # fs) x = app fs (f x)" |
|
26265 | 204 |
|
28315 | 205 |
lemma "app (fs @ gs) x = app gs (app fs x)" |
206 |
quickcheck [generator = code] |
|
207 |
by (induct fs arbitrary: x) simp_all |
|
208 |
||
209 |
lemma "app (fs @ gs) x = app fs (app gs x)" |
|
210 |
quickcheck [generator = code] |
|
211 |
oops |
|
212 |
||
213 |
primrec occurs :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where |
|
214 |
"occurs a [] = 0" |
|
215 |
| "occurs a (x#xs) = (if (x=a) then Suc(occurs a xs) else occurs a xs)" |
|
26265 | 216 |
|
28315 | 217 |
primrec del1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
218 |
"del1 a [] = []" |
|
219 |
| "del1 a (x#xs) = (if (x=a) then xs else (x#del1 a xs))" |
|
220 |
||
221 |
lemma "Suc (occurs a (del1 a xs)) = occurs a xs" |
|
222 |
-- {* Wrong. Precondition needed.*} |
|
223 |
quickcheck [generator = code] |
|
224 |
oops |
|
26265 | 225 |
|
28315 | 226 |
lemma "xs ~= [] \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs" |
227 |
quickcheck [generator = code] |
|
228 |
-- {* Also wrong.*} |
|
229 |
oops |
|
230 |
||
231 |
lemma "0 < occurs a xs \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs" |
|
232 |
quickcheck [generator = code] |
|
233 |
by (induct xs) auto |
|
26265 | 234 |
|
28315 | 235 |
primrec replace :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
236 |
"replace a b [] = []" |
|
237 |
| "replace a b (x#xs) = (if (x=a) then (b#(replace a b xs)) |
|
238 |
else (x#(replace a b xs)))" |
|
239 |
||
240 |
lemma "occurs a xs = occurs b (replace a b xs)" |
|
241 |
quickcheck [generator = code] |
|
242 |
-- {* Wrong. Precondition needed.*} |
|
243 |
oops |
|
244 |
||
245 |
lemma "occurs b xs = 0 \<or> a=b \<longrightarrow> occurs a xs = occurs b (replace a b xs)" |
|
246 |
quickcheck [generator = code] |
|
247 |
by (induct xs) simp_all |
|
248 |
||
249 |
||
250 |
subsection {* Trees *} |
|
251 |
||
252 |
datatype 'a tree = Twig | Leaf 'a | Branch "'a tree" "'a tree" |
|
253 |
||
254 |
primrec leaves :: "'a tree \<Rightarrow> 'a list" where |
|
255 |
"leaves Twig = []" |
|
256 |
| "leaves (Leaf a) = [a]" |
|
257 |
| "leaves (Branch l r) = (leaves l) @ (leaves r)" |
|
258 |
||
259 |
primrec plant :: "'a list \<Rightarrow> 'a tree" where |
|
260 |
"plant [] = Twig " |
|
261 |
| "plant (x#xs) = Branch (Leaf x) (plant xs)" |
|
26265 | 262 |
|
28315 | 263 |
primrec mirror :: "'a tree \<Rightarrow> 'a tree" where |
264 |
"mirror (Twig) = Twig " |
|
265 |
| "mirror (Leaf a) = Leaf a " |
|
266 |
| "mirror (Branch l r) = Branch (mirror r) (mirror l)" |
|
26265 | 267 |
|
28315 | 268 |
theorem "plant (rev (leaves xt)) = mirror xt" |
269 |
quickcheck [generator = code] |
|
270 |
--{* Wrong! *} |
|
271 |
oops |
|
272 |
||
273 |
theorem "plant (leaves xt @ leaves yt) = Branch xt yt" |
|
274 |
quickcheck [generator = code] |
|
275 |
--{* Wrong! *} |
|
276 |
oops |
|
277 |
||
278 |
datatype 'a ntree = Tip "'a" | Node "'a" "'a ntree" "'a ntree" |
|
26265 | 279 |
|
28315 | 280 |
primrec inOrder :: "'a ntree \<Rightarrow> 'a list" where |
281 |
"inOrder (Tip a)= [a]" |
|
282 |
| "inOrder (Node f x y) = (inOrder x)@[f]@(inOrder y)" |
|
283 |
||
284 |
primrec root :: "'a ntree \<Rightarrow> 'a" where |
|
285 |
"root (Tip a) = a" |
|
286 |
| "root (Node f x y) = f" |
|
26265 | 287 |
|
28315 | 288 |
theorem "hd (inOrder xt) = root xt" |
289 |
quickcheck [generator = code] |
|
290 |
--{* Wrong! *} |
|
291 |
oops |
|
26325 | 292 |
|
28315 | 293 |
lemma "int (f k) = k" |
294 |
quickcheck [generator = code] |
|
295 |
oops |
|
26325 | 296 |
|
26265 | 297 |
end |