author | blanchet |
Wed, 15 Dec 2010 11:26:28 +0100 | |
changeset 41140 | 9c68004b8c9d |
parent 40917 | c288fd2ead5a |
child 41231 | 2e901158675e |
permissions | -rw-r--r-- |
14592 | 1 |
(* Title: HOL/ex/Quickcheck_Examples.thy |
40645
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
2 |
Author: Stefan Berghofer, Lukas Bulwahn |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
3 |
Copyright 2004 - 2010 TU Muenchen |
14592 | 4 |
*) |
5 |
||
6 |
header {* Examples for the 'quickcheck' command *} |
|
7 |
||
28314 | 8 |
theory Quickcheck_Examples |
9 |
imports Main |
|
10 |
begin |
|
14592 | 11 |
|
12 |
text {* |
|
13 |
The 'quickcheck' command allows to find counterexamples by evaluating |
|
40654 | 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 |
||
14592 | 22 |
*} |
23 |
||
24 |
subsection {* Lists *} |
|
25 |
||
26 |
theorem "map g (map f xs) = map (g o f) xs" |
|
40917 | 27 |
quickcheck[random, expect = no_counterexample] |
28 |
quickcheck[exhaustive, size = 3, expect = no_counterexample] |
|
14592 | 29 |
oops |
30 |
||
31 |
theorem "map g (map f xs) = map (f o g) xs" |
|
40917 | 32 |
quickcheck[random, expect = counterexample] |
33 |
quickcheck[exhaustive, expect = counterexample] |
|
14592 | 34 |
oops |
35 |
||
36 |
theorem "rev (xs @ ys) = rev ys @ rev xs" |
|
40917 | 37 |
quickcheck[random, expect = no_counterexample] |
38 |
quickcheck[exhaustive, expect = no_counterexample] |
|
14592 | 39 |
oops |
40 |
||
41 |
theorem "rev (xs @ ys) = rev xs @ rev ys" |
|
40917 | 42 |
quickcheck[random, expect = counterexample] |
43 |
quickcheck[exhaustive, expect = counterexample] |
|
14592 | 44 |
oops |
45 |
||
46 |
theorem "rev (rev xs) = xs" |
|
40917 | 47 |
quickcheck[random, expect = no_counterexample] |
48 |
quickcheck[exhaustive, expect = no_counterexample] |
|
14592 | 49 |
oops |
50 |
||
51 |
theorem "rev xs = xs" |
|
40917 | 52 |
quickcheck[random, expect = counterexample] |
53 |
quickcheck[exhaustive, expect = counterexample] |
|
14592 | 54 |
oops |
55 |
||
25891 | 56 |
text {* An example involving functions inside other data structures *} |
57 |
||
28314 | 58 |
primrec app :: "('a \<Rightarrow> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a" where |
25891 | 59 |
"app [] x = x" |
28314 | 60 |
| "app (f # fs) x = app fs (f x)" |
25891 | 61 |
|
62 |
lemma "app (fs @ gs) x = app gs (app fs x)" |
|
40917 | 63 |
quickcheck[random, expect = no_counterexample] |
64 |
quickcheck[exhaustive, size = 4, expect = no_counterexample] |
|
25891 | 65 |
by (induct fs arbitrary: x) simp_all |
66 |
||
67 |
lemma "app (fs @ gs) x = app fs (app gs x)" |
|
40917 | 68 |
quickcheck[random, expect = counterexample] |
69 |
quickcheck[exhaustive, expect = counterexample] |
|
25891 | 70 |
oops |
71 |
||
28314 | 72 |
primrec occurs :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where |
14592 | 73 |
"occurs a [] = 0" |
28314 | 74 |
| "occurs a (x#xs) = (if (x=a) then Suc(occurs a xs) else occurs a xs)" |
14592 | 75 |
|
28314 | 76 |
primrec del1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
14592 | 77 |
"del1 a [] = []" |
28314 | 78 |
| "del1 a (x#xs) = (if (x=a) then xs else (x#del1 a xs))" |
14592 | 79 |
|
25891 | 80 |
text {* A lemma, you'd think to be true from our experience with delAll *} |
14592 | 81 |
lemma "Suc (occurs a (del1 a xs)) = occurs a xs" |
82 |
-- {* Wrong. Precondition needed.*} |
|
40917 | 83 |
quickcheck[random, expect = counterexample] |
84 |
quickcheck[exhaustive, expect = counterexample] |
|
14592 | 85 |
oops |
86 |
||
87 |
lemma "xs ~= [] \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs" |
|
40917 | 88 |
quickcheck[random, expect = counterexample] |
89 |
quickcheck[exhaustive, expect = counterexample] |
|
14592 | 90 |
-- {* Also wrong.*} |
91 |
oops |
|
92 |
||
93 |
lemma "0 < occurs a xs \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs" |
|
40917 | 94 |
quickcheck[random, expect = no_counterexample] |
95 |
quickcheck[exhaustive, expect = no_counterexample] |
|
28314 | 96 |
by (induct xs) auto |
14592 | 97 |
|
28314 | 98 |
primrec replace :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
14592 | 99 |
"replace a b [] = []" |
28314 | 100 |
| "replace a b (x#xs) = (if (x=a) then (b#(replace a b xs)) |
14592 | 101 |
else (x#(replace a b xs)))" |
102 |
||
103 |
lemma "occurs a xs = occurs b (replace a b xs)" |
|
40917 | 104 |
quickcheck[random, expect = counterexample] |
105 |
quickcheck[exhaustive, expect = counterexample] |
|
14592 | 106 |
-- {* Wrong. Precondition needed.*} |
107 |
oops |
|
108 |
||
109 |
lemma "occurs b xs = 0 \<or> a=b \<longrightarrow> occurs a xs = occurs b (replace a b xs)" |
|
40917 | 110 |
quickcheck[random, expect = no_counterexample] |
111 |
quickcheck[exhaustive, expect = no_counterexample] |
|
28314 | 112 |
by (induct xs) simp_all |
14592 | 113 |
|
114 |
||
115 |
subsection {* Trees *} |
|
116 |
||
117 |
datatype 'a tree = Twig | Leaf 'a | Branch "'a tree" "'a tree" |
|
118 |
||
28314 | 119 |
primrec leaves :: "'a tree \<Rightarrow> 'a list" where |
14592 | 120 |
"leaves Twig = []" |
28314 | 121 |
| "leaves (Leaf a) = [a]" |
122 |
| "leaves (Branch l r) = (leaves l) @ (leaves r)" |
|
14592 | 123 |
|
28314 | 124 |
primrec plant :: "'a list \<Rightarrow> 'a tree" where |
14592 | 125 |
"plant [] = Twig " |
28314 | 126 |
| "plant (x#xs) = Branch (Leaf x) (plant xs)" |
14592 | 127 |
|
28314 | 128 |
primrec mirror :: "'a tree \<Rightarrow> 'a tree" where |
14592 | 129 |
"mirror (Twig) = Twig " |
28314 | 130 |
| "mirror (Leaf a) = Leaf a " |
131 |
| "mirror (Branch l r) = Branch (mirror r) (mirror l)" |
|
14592 | 132 |
|
133 |
theorem "plant (rev (leaves xt)) = mirror xt" |
|
40917 | 134 |
quickcheck[random, expect = counterexample] |
135 |
quickcheck[exhaustive, expect = counterexample] |
|
14592 | 136 |
--{* Wrong! *} |
137 |
oops |
|
138 |
||
139 |
theorem "plant((leaves xt) @ (leaves yt)) = Branch xt yt" |
|
40917 | 140 |
quickcheck[random, expect = counterexample] |
141 |
quickcheck[exhaustive, expect = counterexample] |
|
14592 | 142 |
--{* Wrong! *} |
143 |
oops |
|
144 |
||
145 |
datatype 'a ntree = Tip "'a" | Node "'a" "'a ntree" "'a ntree" |
|
146 |
||
28314 | 147 |
primrec inOrder :: "'a ntree \<Rightarrow> 'a list" where |
14592 | 148 |
"inOrder (Tip a)= [a]" |
28314 | 149 |
| "inOrder (Node f x y) = (inOrder x)@[f]@(inOrder y)" |
14592 | 150 |
|
28314 | 151 |
primrec root :: "'a ntree \<Rightarrow> 'a" where |
14592 | 152 |
"root (Tip a) = a" |
28314 | 153 |
| "root (Node f x y) = f" |
14592 | 154 |
|
28314 | 155 |
theorem "hd (inOrder xt) = root xt" |
40917 | 156 |
quickcheck[random, expect = counterexample] |
157 |
quickcheck[exhaustive, expect = counterexample] |
|
158 |
--{* Wrong! *} |
|
14592 | 159 |
oops |
160 |
||
40645
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
161 |
|
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
162 |
subsection {* Exhaustive Testing beats Random Testing *} |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
163 |
|
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
164 |
text {* Here are some examples from mutants from the List theory |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
165 |
where exhaustive testing beats random testing *} |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
166 |
|
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
167 |
lemma |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
168 |
"[] ~= xs ==> hd xs = last (x # xs)" |
40917 | 169 |
quickcheck[random] |
170 |
quickcheck[exhaustive, expect = counterexample] |
|
40645
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
171 |
oops |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
172 |
|
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
173 |
lemma |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
174 |
assumes "!!i. [| i < n; i < length xs |] ==> P (xs ! i)" "n < length xs ==> ~ P (xs ! n)" |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
175 |
shows "drop n xs = takeWhile P xs" |
40917 | 176 |
quickcheck[random, iterations = 10000, quiet] |
177 |
quickcheck[exhaustive, expect = counterexample] |
|
40645
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
178 |
oops |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
179 |
|
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
180 |
lemma |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
181 |
"i < length (List.transpose (List.transpose xs)) ==> xs ! i = map (%xs. xs ! i) [ys<-xs. i < length ys]" |
40917 | 182 |
quickcheck[random, iterations = 10000] |
183 |
quickcheck[exhaustive, expect = counterexample] |
|
40645
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
184 |
oops |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
185 |
|
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
186 |
lemma |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
187 |
"i < n - m ==> f (lcm m i) = map f [m..<n] ! i" |
40917 | 188 |
quickcheck[random, iterations = 10000, finite_types = false] |
189 |
quickcheck[exhaustive, finite_types = false, expect = counterexample] |
|
40645
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
190 |
oops |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
191 |
|
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
192 |
lemma |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
193 |
"i < n - m ==> f (lcm m i) = map f [m..<n] ! i" |
40917 | 194 |
quickcheck[random, iterations = 10000, finite_types = false] |
195 |
quickcheck[exhaustive, finite_types = false, expect = counterexample] |
|
40645
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
196 |
oops |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
197 |
|
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
198 |
lemma |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
199 |
"ns ! k < length ns ==> k <= listsum ns" |
40917 | 200 |
quickcheck[random, iterations = 10000, finite_types = false, quiet] |
201 |
quickcheck[exhaustive, finite_types = false, expect = counterexample] |
|
40645
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
202 |
oops |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
203 |
|
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
204 |
lemma |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
205 |
"[| ys = x # xs1; zs = xs1 @ xs |] ==> ys @ zs = x # xs" |
40917 | 206 |
quickcheck[random, iterations = 10000] |
207 |
quickcheck[exhaustive, expect = counterexample] |
|
40645
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
208 |
oops |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
209 |
|
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
210 |
lemma |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
211 |
"i < length xs ==> take (Suc i) xs = [] @ xs ! i # take i xs" |
40917 | 212 |
quickcheck[random, iterations = 10000] |
213 |
quickcheck[exhaustive, expect = counterexample] |
|
40645
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
214 |
oops |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
215 |
|
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
216 |
lemma |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
217 |
"i < length xs ==> take (Suc i) xs = (xs ! i # xs) @ take i []" |
40917 | 218 |
quickcheck[random, iterations = 10000] |
219 |
quickcheck[exhaustive, expect = counterexample] |
|
40645
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
220 |
oops |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
221 |
|
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
222 |
lemma |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
223 |
"[| sorted (rev (map length xs)); i < length xs |] ==> xs ! i = map (%ys. ys ! i) [ys<-remdups xs. i < length ys]" |
40917 | 224 |
quickcheck[random] |
225 |
quickcheck[exhaustive, expect = counterexample] |
|
40645
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
226 |
oops |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
227 |
|
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
228 |
lemma |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
229 |
"[| sorted (rev (map length xs)); i < length xs |] ==> xs ! i = map (%ys. ys ! i) [ys<-List.transpose xs. {..<i} (length ys)]" |
40917 | 230 |
quickcheck[random] |
231 |
quickcheck[exhaustive, expect = counterexample] |
|
40645
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
232 |
oops |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
233 |
|
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
234 |
lemma |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
235 |
"(ys = zs) = (xs @ ys = splice xs zs)" |
40917 | 236 |
quickcheck[random] |
237 |
quickcheck[exhaustive, expect = counterexample] |
|
40645
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
238 |
oops |
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
bulwahn
parents:
37929
diff
changeset
|
239 |
|
40654 | 240 |
section {* Examples with quantifiers *} |
241 |
||
242 |
text {* |
|
243 |
These examples show that we can handle quantifiers. |
|
244 |
*} |
|
245 |
||
246 |
lemma "(\<exists>x. P x) \<longrightarrow> (\<forall>x. P x)" |
|
40917 | 247 |
quickcheck[random, expect = counterexample] |
248 |
quickcheck[exhaustive, expect = counterexample] |
|
40654 | 249 |
oops |
250 |
||
251 |
lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>y. \<forall>x. P x y)" |
|
40917 | 252 |
quickcheck[random, expect = counterexample] |
40654 | 253 |
quickcheck[expect = counterexample] |
254 |
oops |
|
255 |
||
256 |
lemma "(\<exists>x. P x) \<longrightarrow> (EX! x. P x)" |
|
40917 | 257 |
quickcheck[random, expect = counterexample] |
40654 | 258 |
quickcheck[expect = counterexample] |
259 |
oops |
|
260 |
||
14592 | 261 |
end |