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