|
1 (* Title: HOL/ex/Quickcheck_Examples.thy |
|
2 Author: Stefan Berghofer, Lukas Bulwahn |
|
3 Copyright 2004 - 2010 TU Muenchen |
|
4 *) |
|
5 |
|
6 header {* Examples for the 'quickcheck' command *} |
|
7 |
|
8 theory Quickcheck_Examples |
|
9 imports Complex_Main "~~/src/HOL/Library/Dlist" "~~/src/HOL/Library/Multiset" |
|
10 begin |
|
11 |
|
12 text {* |
|
13 The 'quickcheck' command allows to find counterexamples by evaluating |
|
14 formulae. |
|
15 Currently, there are two different exploration schemes: |
|
16 - random testing: this is incomplete, but explores the search space faster. |
|
17 - exhaustive testing: this is complete, but increasing the depth leads to |
|
18 exponentially many assignments. |
|
19 |
|
20 quickcheck can handle quantifiers on finite universes. |
|
21 |
|
22 *} |
|
23 |
|
24 declare [[quickcheck_timeout = 3600]] |
|
25 |
|
26 subsection {* Lists *} |
|
27 |
|
28 theorem "map g (map f xs) = map (g o f) xs" |
|
29 quickcheck[random, expect = no_counterexample] |
|
30 quickcheck[exhaustive, size = 3, expect = no_counterexample] |
|
31 oops |
|
32 |
|
33 theorem "map g (map f xs) = map (f o g) xs" |
|
34 quickcheck[random, expect = counterexample] |
|
35 quickcheck[exhaustive, expect = counterexample] |
|
36 oops |
|
37 |
|
38 theorem "rev (xs @ ys) = rev ys @ rev xs" |
|
39 quickcheck[random, expect = no_counterexample] |
|
40 quickcheck[exhaustive, expect = no_counterexample] |
|
41 quickcheck[exhaustive, size = 1000, timeout = 0.1] |
|
42 oops |
|
43 |
|
44 theorem "rev (xs @ ys) = rev xs @ rev ys" |
|
45 quickcheck[random, expect = counterexample] |
|
46 quickcheck[exhaustive, expect = counterexample] |
|
47 oops |
|
48 |
|
49 theorem "rev (rev xs) = xs" |
|
50 quickcheck[random, expect = no_counterexample] |
|
51 quickcheck[exhaustive, expect = no_counterexample] |
|
52 oops |
|
53 |
|
54 theorem "rev xs = xs" |
|
55 quickcheck[tester = random, finite_types = true, report = false, expect = counterexample] |
|
56 quickcheck[tester = random, finite_types = false, report = false, expect = counterexample] |
|
57 quickcheck[tester = random, finite_types = true, report = true, expect = counterexample] |
|
58 quickcheck[tester = random, finite_types = false, report = true, expect = counterexample] |
|
59 quickcheck[tester = exhaustive, finite_types = true, expect = counterexample] |
|
60 quickcheck[tester = exhaustive, finite_types = false, expect = counterexample] |
|
61 oops |
|
62 |
|
63 |
|
64 text {* An example involving functions inside other data structures *} |
|
65 |
|
66 primrec app :: "('a \<Rightarrow> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a" where |
|
67 "app [] x = x" |
|
68 | "app (f # fs) x = app fs (f x)" |
|
69 |
|
70 lemma "app (fs @ gs) x = app gs (app fs x)" |
|
71 quickcheck[random, expect = no_counterexample] |
|
72 quickcheck[exhaustive, size = 4, expect = no_counterexample] |
|
73 by (induct fs arbitrary: x) simp_all |
|
74 |
|
75 lemma "app (fs @ gs) x = app fs (app gs x)" |
|
76 quickcheck[random, expect = counterexample] |
|
77 quickcheck[exhaustive, expect = counterexample] |
|
78 oops |
|
79 |
|
80 primrec occurs :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where |
|
81 "occurs a [] = 0" |
|
82 | "occurs a (x#xs) = (if (x=a) then Suc(occurs a xs) else occurs a xs)" |
|
83 |
|
84 primrec del1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
85 "del1 a [] = []" |
|
86 | "del1 a (x#xs) = (if (x=a) then xs else (x#del1 a xs))" |
|
87 |
|
88 text {* A lemma, you'd think to be true from our experience with delAll *} |
|
89 lemma "Suc (occurs a (del1 a xs)) = occurs a xs" |
|
90 -- {* Wrong. Precondition needed.*} |
|
91 quickcheck[random, expect = counterexample] |
|
92 quickcheck[exhaustive, expect = counterexample] |
|
93 oops |
|
94 |
|
95 lemma "xs ~= [] \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs" |
|
96 quickcheck[random, expect = counterexample] |
|
97 quickcheck[exhaustive, expect = counterexample] |
|
98 -- {* Also wrong.*} |
|
99 oops |
|
100 |
|
101 lemma "0 < occurs a xs \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs" |
|
102 quickcheck[random, expect = no_counterexample] |
|
103 quickcheck[exhaustive, expect = no_counterexample] |
|
104 by (induct xs) auto |
|
105 |
|
106 primrec replace :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
107 "replace a b [] = []" |
|
108 | "replace a b (x#xs) = (if (x=a) then (b#(replace a b xs)) |
|
109 else (x#(replace a b xs)))" |
|
110 |
|
111 lemma "occurs a xs = occurs b (replace a b xs)" |
|
112 quickcheck[random, expect = counterexample] |
|
113 quickcheck[exhaustive, expect = counterexample] |
|
114 -- {* Wrong. Precondition needed.*} |
|
115 oops |
|
116 |
|
117 lemma "occurs b xs = 0 \<or> a=b \<longrightarrow> occurs a xs = occurs b (replace a b xs)" |
|
118 quickcheck[random, expect = no_counterexample] |
|
119 quickcheck[exhaustive, expect = no_counterexample] |
|
120 by (induct xs) simp_all |
|
121 |
|
122 |
|
123 subsection {* Trees *} |
|
124 |
|
125 datatype 'a tree = Twig | Leaf 'a | Branch "'a tree" "'a tree" |
|
126 |
|
127 primrec leaves :: "'a tree \<Rightarrow> 'a list" where |
|
128 "leaves Twig = []" |
|
129 | "leaves (Leaf a) = [a]" |
|
130 | "leaves (Branch l r) = (leaves l) @ (leaves r)" |
|
131 |
|
132 primrec plant :: "'a list \<Rightarrow> 'a tree" where |
|
133 "plant [] = Twig " |
|
134 | "plant (x#xs) = Branch (Leaf x) (plant xs)" |
|
135 |
|
136 primrec mirror :: "'a tree \<Rightarrow> 'a tree" where |
|
137 "mirror (Twig) = Twig " |
|
138 | "mirror (Leaf a) = Leaf a " |
|
139 | "mirror (Branch l r) = Branch (mirror r) (mirror l)" |
|
140 |
|
141 theorem "plant (rev (leaves xt)) = mirror xt" |
|
142 quickcheck[random, expect = counterexample] |
|
143 quickcheck[exhaustive, expect = counterexample] |
|
144 --{* Wrong! *} |
|
145 oops |
|
146 |
|
147 theorem "plant((leaves xt) @ (leaves yt)) = Branch xt yt" |
|
148 quickcheck[random, expect = counterexample] |
|
149 quickcheck[exhaustive, expect = counterexample] |
|
150 --{* Wrong! *} |
|
151 oops |
|
152 |
|
153 datatype 'a ntree = Tip "'a" | Node "'a" "'a ntree" "'a ntree" |
|
154 |
|
155 primrec inOrder :: "'a ntree \<Rightarrow> 'a list" where |
|
156 "inOrder (Tip a)= [a]" |
|
157 | "inOrder (Node f x y) = (inOrder x)@[f]@(inOrder y)" |
|
158 |
|
159 primrec root :: "'a ntree \<Rightarrow> 'a" where |
|
160 "root (Tip a) = a" |
|
161 | "root (Node f x y) = f" |
|
162 |
|
163 theorem "hd (inOrder xt) = root xt" |
|
164 quickcheck[random, expect = counterexample] |
|
165 quickcheck[exhaustive, expect = counterexample] |
|
166 --{* Wrong! *} |
|
167 oops |
|
168 |
|
169 |
|
170 subsection {* Exhaustive Testing beats Random Testing *} |
|
171 |
|
172 text {* Here are some examples from mutants from the List theory |
|
173 where exhaustive testing beats random testing *} |
|
174 |
|
175 lemma |
|
176 "[] ~= xs ==> hd xs = last (x # xs)" |
|
177 quickcheck[random] |
|
178 quickcheck[exhaustive, expect = counterexample] |
|
179 oops |
|
180 |
|
181 lemma |
|
182 assumes "!!i. [| i < n; i < length xs |] ==> P (xs ! i)" "n < length xs ==> ~ P (xs ! n)" |
|
183 shows "drop n xs = takeWhile P xs" |
|
184 quickcheck[random, iterations = 10000, quiet] |
|
185 quickcheck[exhaustive, expect = counterexample] |
|
186 oops |
|
187 |
|
188 lemma |
|
189 "i < length (List.transpose (List.transpose xs)) ==> xs ! i = map (%xs. xs ! i) [ys<-xs. i < length ys]" |
|
190 quickcheck[random, iterations = 10000] |
|
191 quickcheck[exhaustive, expect = counterexample] |
|
192 oops |
|
193 |
|
194 lemma |
|
195 "i < n - m ==> f (lcm m i) = map f [m..<n] ! i" |
|
196 quickcheck[random, iterations = 10000, finite_types = false] |
|
197 quickcheck[exhaustive, finite_types = false, expect = counterexample] |
|
198 oops |
|
199 |
|
200 lemma |
|
201 "i < n - m ==> f (lcm m i) = map f [m..<n] ! i" |
|
202 quickcheck[random, iterations = 10000, finite_types = false] |
|
203 quickcheck[exhaustive, finite_types = false, expect = counterexample] |
|
204 oops |
|
205 |
|
206 lemma |
|
207 "ns ! k < length ns ==> k <= listsum ns" |
|
208 quickcheck[random, iterations = 10000, finite_types = false, quiet] |
|
209 quickcheck[exhaustive, finite_types = false, expect = counterexample] |
|
210 oops |
|
211 |
|
212 lemma |
|
213 "[| ys = x # xs1; zs = xs1 @ xs |] ==> ys @ zs = x # xs" |
|
214 quickcheck[random, iterations = 10000] |
|
215 quickcheck[exhaustive, expect = counterexample] |
|
216 oops |
|
217 |
|
218 lemma |
|
219 "i < length xs ==> take (Suc i) xs = [] @ xs ! i # take i xs" |
|
220 quickcheck[random, iterations = 10000] |
|
221 quickcheck[exhaustive, expect = counterexample] |
|
222 oops |
|
223 |
|
224 lemma |
|
225 "i < length xs ==> take (Suc i) xs = (xs ! i # xs) @ take i []" |
|
226 quickcheck[random, iterations = 10000] |
|
227 quickcheck[exhaustive, expect = counterexample] |
|
228 oops |
|
229 |
|
230 lemma |
|
231 "[| sorted (rev (map length xs)); i < length xs |] ==> xs ! i = map (%ys. ys ! i) [ys<-remdups xs. i < length ys]" |
|
232 quickcheck[random] |
|
233 quickcheck[exhaustive, expect = counterexample] |
|
234 oops |
|
235 |
|
236 lemma |
|
237 "[| sorted (rev (map length xs)); i < length xs |] ==> xs ! i = map (%ys. ys ! i) [ys<-List.transpose xs. length ys \<in> {..<i}]" |
|
238 quickcheck[random] |
|
239 quickcheck[exhaustive, expect = counterexample] |
|
240 oops |
|
241 |
|
242 lemma |
|
243 "(ys = zs) = (xs @ ys = splice xs zs)" |
|
244 quickcheck[random] |
|
245 quickcheck[exhaustive, expect = counterexample] |
|
246 oops |
|
247 |
|
248 subsection {* Examples with quantifiers *} |
|
249 |
|
250 text {* |
|
251 These examples show that we can handle quantifiers. |
|
252 *} |
|
253 |
|
254 lemma "(\<exists>x. P x) \<longrightarrow> (\<forall>x. P x)" |
|
255 quickcheck[random, expect = counterexample] |
|
256 quickcheck[exhaustive, expect = counterexample] |
|
257 oops |
|
258 |
|
259 lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>y. \<forall>x. P x y)" |
|
260 quickcheck[random, expect = counterexample] |
|
261 quickcheck[expect = counterexample] |
|
262 oops |
|
263 |
|
264 lemma "(\<exists>x. P x) \<longrightarrow> (EX! x. P x)" |
|
265 quickcheck[random, expect = counterexample] |
|
266 quickcheck[expect = counterexample] |
|
267 oops |
|
268 |
|
269 |
|
270 subsection {* Examples with sets *} |
|
271 |
|
272 lemma |
|
273 "{} = A Un - A" |
|
274 quickcheck[exhaustive, expect = counterexample] |
|
275 oops |
|
276 |
|
277 lemma |
|
278 "[| bij_betw f A B; bij_betw f C D |] ==> bij_betw f (A Un C) (B Un D)" |
|
279 quickcheck[exhaustive, expect = counterexample] |
|
280 oops |
|
281 |
|
282 |
|
283 subsection {* Examples with relations *} |
|
284 |
|
285 lemma |
|
286 "acyclic (R :: ('a * 'a) set) ==> acyclic S ==> acyclic (R Un S)" |
|
287 quickcheck[exhaustive, expect = counterexample] |
|
288 oops |
|
289 |
|
290 lemma |
|
291 "acyclic (R :: (nat * nat) set) ==> acyclic S ==> acyclic (R Un S)" |
|
292 quickcheck[exhaustive, expect = counterexample] |
|
293 oops |
|
294 |
|
295 (* FIXME: some dramatic performance decrease after changing the code equation of the ntrancl *) |
|
296 lemma |
|
297 "(x, z) : rtrancl (R Un S) ==> \<exists> y. (x, y) : rtrancl R & (y, z) : rtrancl S" |
|
298 (*quickcheck[exhaustive, expect = counterexample]*) |
|
299 oops |
|
300 |
|
301 lemma |
|
302 "wf (R :: ('a * 'a) set) ==> wf S ==> wf (R Un S)" |
|
303 quickcheck[exhaustive, expect = counterexample] |
|
304 oops |
|
305 |
|
306 lemma |
|
307 "wf (R :: (nat * nat) set) ==> wf S ==> wf (R Un S)" |
|
308 quickcheck[exhaustive, expect = counterexample] |
|
309 oops |
|
310 |
|
311 lemma |
|
312 "wf (R :: (int * int) set) ==> wf S ==> wf (R Un S)" |
|
313 quickcheck[exhaustive, expect = counterexample] |
|
314 oops |
|
315 |
|
316 |
|
317 subsection {* Examples with the descriptive operator *} |
|
318 |
|
319 lemma |
|
320 "(THE x. x = a) = b" |
|
321 quickcheck[random, expect = counterexample] |
|
322 quickcheck[exhaustive, expect = counterexample] |
|
323 oops |
|
324 |
|
325 subsection {* Examples with Multisets *} |
|
326 |
|
327 lemma |
|
328 "X + Y = Y + (Z :: 'a multiset)" |
|
329 quickcheck[random, expect = counterexample] |
|
330 quickcheck[exhaustive, expect = counterexample] |
|
331 oops |
|
332 |
|
333 lemma |
|
334 "X - Y = Y - (Z :: 'a multiset)" |
|
335 quickcheck[random, expect = counterexample] |
|
336 quickcheck[exhaustive, expect = counterexample] |
|
337 oops |
|
338 |
|
339 lemma |
|
340 "N + M - N = (N::'a multiset)" |
|
341 quickcheck[random, expect = counterexample] |
|
342 quickcheck[exhaustive, expect = counterexample] |
|
343 oops |
|
344 |
|
345 subsection {* Examples with numerical types *} |
|
346 |
|
347 text {* |
|
348 Quickcheck supports the common types nat, int, rat and real. |
|
349 *} |
|
350 |
|
351 lemma |
|
352 "(x :: nat) > 0 ==> y > 0 ==> z > 0 ==> x * x + y * y \<noteq> z * z" |
|
353 quickcheck[exhaustive, size = 10, expect = counterexample] |
|
354 quickcheck[random, size = 10] |
|
355 oops |
|
356 |
|
357 lemma |
|
358 "(x :: int) > 0 ==> y > 0 ==> z > 0 ==> x * x + y * y \<noteq> z * z" |
|
359 quickcheck[exhaustive, size = 10, expect = counterexample] |
|
360 quickcheck[random, size = 10] |
|
361 oops |
|
362 |
|
363 lemma |
|
364 "(x :: rat) > 0 ==> y > 0 ==> z > 0 ==> x * x + y * y \<noteq> z * z" |
|
365 quickcheck[exhaustive, size = 10, expect = counterexample] |
|
366 quickcheck[random, size = 10] |
|
367 oops |
|
368 |
|
369 lemma "(x :: rat) >= 0" |
|
370 quickcheck[random, expect = counterexample] |
|
371 quickcheck[exhaustive, expect = counterexample] |
|
372 oops |
|
373 |
|
374 lemma |
|
375 "(x :: real) > 0 ==> y > 0 ==> z > 0 ==> x * x + y * y \<noteq> z * z" |
|
376 quickcheck[exhaustive, size = 10, expect = counterexample] |
|
377 quickcheck[random, size = 10] |
|
378 oops |
|
379 |
|
380 lemma "(x :: real) >= 0" |
|
381 quickcheck[random, expect = counterexample] |
|
382 quickcheck[exhaustive, expect = counterexample] |
|
383 oops |
|
384 |
|
385 subsubsection {* floor and ceiling functions *} |
|
386 |
|
387 lemma |
|
388 "floor x + floor y = floor (x + y :: rat)" |
|
389 quickcheck[expect = counterexample] |
|
390 oops |
|
391 |
|
392 lemma |
|
393 "floor x + floor y = floor (x + y :: real)" |
|
394 quickcheck[expect = counterexample] |
|
395 oops |
|
396 |
|
397 lemma |
|
398 "ceiling x + ceiling y = ceiling (x + y :: rat)" |
|
399 quickcheck[expect = counterexample] |
|
400 oops |
|
401 |
|
402 lemma |
|
403 "ceiling x + ceiling y = ceiling (x + y :: real)" |
|
404 quickcheck[expect = counterexample] |
|
405 oops |
|
406 |
|
407 subsection {* Examples with abstract types *} |
|
408 |
|
409 lemma |
|
410 "Dlist.length (Dlist.remove x xs) = Dlist.length xs - 1" |
|
411 quickcheck[exhaustive] |
|
412 quickcheck[random] |
|
413 oops |
|
414 |
|
415 lemma |
|
416 "Dlist.length (Dlist.insert x xs) = Dlist.length xs + 1" |
|
417 quickcheck[exhaustive] |
|
418 quickcheck[random] |
|
419 oops |
|
420 |
|
421 subsection {* Examples with Records *} |
|
422 |
|
423 record point = |
|
424 xpos :: nat |
|
425 ypos :: nat |
|
426 |
|
427 lemma |
|
428 "xpos r = xpos r' ==> r = r'" |
|
429 quickcheck[exhaustive, expect = counterexample] |
|
430 quickcheck[random, expect = counterexample] |
|
431 oops |
|
432 |
|
433 datatype colour = Red | Green | Blue |
|
434 |
|
435 record cpoint = point + |
|
436 colour :: colour |
|
437 |
|
438 lemma |
|
439 "xpos r = xpos r' ==> ypos r = ypos r' ==> (r :: cpoint) = r'" |
|
440 quickcheck[exhaustive, expect = counterexample] |
|
441 quickcheck[random, expect = counterexample] |
|
442 oops |
|
443 |
|
444 subsection {* Examples with locales *} |
|
445 |
|
446 locale Truth |
|
447 |
|
448 context Truth |
|
449 begin |
|
450 |
|
451 lemma "False" |
|
452 quickcheck[exhaustive, expect = counterexample] |
|
453 oops |
|
454 |
|
455 end |
|
456 |
|
457 interpretation Truth . |
|
458 |
|
459 context Truth |
|
460 begin |
|
461 |
|
462 lemma "False" |
|
463 quickcheck[exhaustive, expect = counterexample] |
|
464 oops |
|
465 |
|
466 end |
|
467 |
|
468 locale antisym = |
|
469 fixes R |
|
470 assumes "R x y --> R y x --> x = y" |
|
471 begin |
|
472 |
|
473 lemma |
|
474 "R x y --> R y z --> R x z" |
|
475 quickcheck[exhaustive, finite_type_size = 2, expect = no_counterexample] |
|
476 quickcheck[exhaustive, expect = counterexample] |
|
477 oops |
|
478 |
|
479 end |
|
480 |
|
481 subsection {* Examples with HOL quantifiers *} |
|
482 |
|
483 lemma |
|
484 "\<forall> xs ys. xs = [] --> xs = ys" |
|
485 quickcheck[exhaustive, expect = counterexample] |
|
486 oops |
|
487 |
|
488 lemma |
|
489 "ys = [] --> (\<forall>xs. xs = [] --> xs = y # ys)" |
|
490 quickcheck[exhaustive, expect = counterexample] |
|
491 oops |
|
492 |
|
493 lemma |
|
494 "\<forall>xs. (\<exists> ys. ys = []) --> xs = ys" |
|
495 quickcheck[exhaustive, expect = counterexample] |
|
496 oops |
|
497 |
|
498 subsection {* Examples with underspecified/partial functions *} |
|
499 |
|
500 lemma |
|
501 "xs = [] ==> hd xs \<noteq> x" |
|
502 quickcheck[exhaustive, expect = no_counterexample] |
|
503 quickcheck[random, report = false, expect = no_counterexample] |
|
504 quickcheck[random, report = true, expect = no_counterexample] |
|
505 oops |
|
506 |
|
507 lemma |
|
508 "xs = [] ==> hd xs = x" |
|
509 quickcheck[exhaustive, expect = no_counterexample] |
|
510 quickcheck[random, report = false, expect = no_counterexample] |
|
511 quickcheck[random, report = true, expect = no_counterexample] |
|
512 oops |
|
513 |
|
514 lemma "xs = [] ==> hd xs = x ==> x = y" |
|
515 quickcheck[exhaustive, expect = no_counterexample] |
|
516 quickcheck[random, report = false, expect = no_counterexample] |
|
517 quickcheck[random, report = true, expect = no_counterexample] |
|
518 oops |
|
519 |
|
520 text {* with the simple testing scheme *} |
|
521 |
|
522 setup {* Exhaustive_Generators.setup_exhaustive_datatype_interpretation *} |
|
523 declare [[quickcheck_full_support = false]] |
|
524 |
|
525 lemma |
|
526 "xs = [] ==> hd xs \<noteq> x" |
|
527 quickcheck[exhaustive, expect = no_counterexample] |
|
528 oops |
|
529 |
|
530 lemma |
|
531 "xs = [] ==> hd xs = x" |
|
532 quickcheck[exhaustive, expect = no_counterexample] |
|
533 oops |
|
534 |
|
535 lemma "xs = [] ==> hd xs = x ==> x = y" |
|
536 quickcheck[exhaustive, expect = no_counterexample] |
|
537 oops |
|
538 |
|
539 declare [[quickcheck_full_support = true]] |
|
540 |
|
541 end |