| author | blanchet | 
| Sun, 01 May 2011 18:37:24 +0200 | |
| changeset 42557 | ae0deb39a254 | 
| parent 42434 | 1914fd5d7c0e | 
| child 42696 | 7c7ca3fc7ce5 | 
| 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: 
37929diff
changeset | 2 | Author: Stefan Berghofer, Lukas Bulwahn | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 3 | Copyright 2004 - 2010 TU Muenchen | 
| 14592 | 4 | *) | 
| 5 | ||
| 6 | header {* Examples for the 'quickcheck' command *}
 | |
| 7 | ||
| 28314 | 8 | theory Quickcheck_Examples | 
| 41231 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 9 | imports Complex_Main | 
| 28314 | 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] | |
| 42087 
5e236f6ef04f
changing timeout behaviour of quickcheck to proceed after command rather than failing; adding a test case for timeout
 bulwahn parents: 
41231diff
changeset | 39 | quickcheck[exhaustive, size = 1000, timeout = 0.1] | 
| 14592 | 40 | oops | 
| 41 | ||
| 42 | theorem "rev (xs @ ys) = rev xs @ rev ys" | |
| 40917 | 43 | quickcheck[random, expect = counterexample] | 
| 44 | quickcheck[exhaustive, expect = counterexample] | |
| 14592 | 45 | oops | 
| 46 | ||
| 47 | theorem "rev (rev xs) = xs" | |
| 40917 | 48 | quickcheck[random, expect = no_counterexample] | 
| 49 | quickcheck[exhaustive, expect = no_counterexample] | |
| 14592 | 50 | oops | 
| 51 | ||
| 52 | theorem "rev xs = xs" | |
| 42159 
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
 bulwahn parents: 
42087diff
changeset | 53 | quickcheck[tester = random, finite_types = true, report = false, expect = counterexample] | 
| 
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
 bulwahn parents: 
42087diff
changeset | 54 | quickcheck[tester = random, finite_types = false, report = false, expect = counterexample] | 
| 
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
 bulwahn parents: 
42087diff
changeset | 55 | quickcheck[tester = random, finite_types = true, report = true, expect = counterexample] | 
| 
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
 bulwahn parents: 
42087diff
changeset | 56 | quickcheck[tester = random, finite_types = false, report = true, expect = counterexample] | 
| 
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
 bulwahn parents: 
42087diff
changeset | 57 | quickcheck[tester = exhaustive, finite_types = true, expect = counterexample] | 
| 
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
 bulwahn parents: 
42087diff
changeset | 58 | quickcheck[tester = exhaustive, finite_types = false, expect = counterexample] | 
| 
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
 bulwahn parents: 
42087diff
changeset | 59 | oops | 
| 
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
 bulwahn parents: 
42087diff
changeset | 60 | |
| 14592 | 61 | |
| 25891 | 62 | text {* An example involving functions inside other data structures *}
 | 
| 63 | ||
| 28314 | 64 | primrec app :: "('a \<Rightarrow> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a" where
 | 
| 25891 | 65 | "app [] x = x" | 
| 28314 | 66 | | "app (f # fs) x = app fs (f x)" | 
| 25891 | 67 | |
| 68 | lemma "app (fs @ gs) x = app gs (app fs x)" | |
| 40917 | 69 | quickcheck[random, expect = no_counterexample] | 
| 70 | quickcheck[exhaustive, size = 4, expect = no_counterexample] | |
| 25891 | 71 | by (induct fs arbitrary: x) simp_all | 
| 72 | ||
| 73 | lemma "app (fs @ gs) x = app fs (app gs x)" | |
| 40917 | 74 | quickcheck[random, expect = counterexample] | 
| 75 | quickcheck[exhaustive, expect = counterexample] | |
| 25891 | 76 | oops | 
| 77 | ||
| 28314 | 78 | primrec occurs :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where | 
| 14592 | 79 | "occurs a [] = 0" | 
| 28314 | 80 | | "occurs a (x#xs) = (if (x=a) then Suc(occurs a xs) else occurs a xs)" | 
| 14592 | 81 | |
| 28314 | 82 | primrec del1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where | 
| 14592 | 83 | "del1 a [] = []" | 
| 28314 | 84 | | "del1 a (x#xs) = (if (x=a) then xs else (x#del1 a xs))" | 
| 14592 | 85 | |
| 25891 | 86 | text {* A lemma, you'd think to be true from our experience with delAll *}
 | 
| 14592 | 87 | lemma "Suc (occurs a (del1 a xs)) = occurs a xs" | 
| 88 |   -- {* Wrong. Precondition needed.*}
 | |
| 40917 | 89 | quickcheck[random, expect = counterexample] | 
| 90 | quickcheck[exhaustive, expect = counterexample] | |
| 14592 | 91 | oops | 
| 92 | ||
| 93 | lemma "xs ~= [] \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs" | |
| 40917 | 94 | quickcheck[random, expect = counterexample] | 
| 95 | quickcheck[exhaustive, expect = counterexample] | |
| 14592 | 96 |     -- {* Also wrong.*}
 | 
| 97 | oops | |
| 98 | ||
| 99 | lemma "0 < occurs a xs \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs" | |
| 40917 | 100 | quickcheck[random, expect = no_counterexample] | 
| 101 | quickcheck[exhaustive, expect = no_counterexample] | |
| 28314 | 102 | by (induct xs) auto | 
| 14592 | 103 | |
| 28314 | 104 | primrec replace :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where | 
| 14592 | 105 | "replace a b [] = []" | 
| 28314 | 106 | | "replace a b (x#xs) = (if (x=a) then (b#(replace a b xs)) | 
| 14592 | 107 | else (x#(replace a b xs)))" | 
| 108 | ||
| 109 | lemma "occurs a xs = occurs b (replace a b xs)" | |
| 40917 | 110 | quickcheck[random, expect = counterexample] | 
| 111 | quickcheck[exhaustive, expect = counterexample] | |
| 14592 | 112 |   -- {* Wrong. Precondition needed.*}
 | 
| 113 | oops | |
| 114 | ||
| 115 | lemma "occurs b xs = 0 \<or> a=b \<longrightarrow> occurs a xs = occurs b (replace a b xs)" | |
| 40917 | 116 | quickcheck[random, expect = no_counterexample] | 
| 117 | quickcheck[exhaustive, expect = no_counterexample] | |
| 28314 | 118 | by (induct xs) simp_all | 
| 14592 | 119 | |
| 120 | ||
| 121 | subsection {* Trees *}
 | |
| 122 | ||
| 123 | datatype 'a tree = Twig | Leaf 'a | Branch "'a tree" "'a tree" | |
| 124 | ||
| 28314 | 125 | primrec leaves :: "'a tree \<Rightarrow> 'a list" where | 
| 14592 | 126 | "leaves Twig = []" | 
| 28314 | 127 | | "leaves (Leaf a) = [a]" | 
| 128 | | "leaves (Branch l r) = (leaves l) @ (leaves r)" | |
| 14592 | 129 | |
| 28314 | 130 | primrec plant :: "'a list \<Rightarrow> 'a tree" where | 
| 14592 | 131 | "plant [] = Twig " | 
| 28314 | 132 | | "plant (x#xs) = Branch (Leaf x) (plant xs)" | 
| 14592 | 133 | |
| 28314 | 134 | primrec mirror :: "'a tree \<Rightarrow> 'a tree" where | 
| 14592 | 135 | "mirror (Twig) = Twig " | 
| 28314 | 136 | | "mirror (Leaf a) = Leaf a " | 
| 137 | | "mirror (Branch l r) = Branch (mirror r) (mirror l)" | |
| 14592 | 138 | |
| 139 | theorem "plant (rev (leaves xt)) = mirror xt" | |
| 40917 | 140 | quickcheck[random, expect = counterexample] | 
| 141 | quickcheck[exhaustive, expect = counterexample] | |
| 14592 | 142 |     --{* Wrong! *} 
 | 
| 143 | oops | |
| 144 | ||
| 145 | theorem "plant((leaves xt) @ (leaves yt)) = Branch xt yt" | |
| 40917 | 146 | quickcheck[random, expect = counterexample] | 
| 147 | quickcheck[exhaustive, expect = counterexample] | |
| 14592 | 148 |     --{* Wrong! *} 
 | 
| 149 | oops | |
| 150 | ||
| 151 | datatype 'a ntree = Tip "'a" | Node "'a" "'a ntree" "'a ntree" | |
| 152 | ||
| 28314 | 153 | primrec inOrder :: "'a ntree \<Rightarrow> 'a list" where | 
| 14592 | 154 | "inOrder (Tip a)= [a]" | 
| 28314 | 155 | | "inOrder (Node f x y) = (inOrder x)@[f]@(inOrder y)" | 
| 14592 | 156 | |
| 28314 | 157 | primrec root :: "'a ntree \<Rightarrow> 'a" where | 
| 14592 | 158 | "root (Tip a) = a" | 
| 28314 | 159 | | "root (Node f x y) = f" | 
| 14592 | 160 | |
| 28314 | 161 | theorem "hd (inOrder xt) = root xt" | 
| 40917 | 162 | quickcheck[random, expect = counterexample] | 
| 163 | quickcheck[exhaustive, expect = counterexample] | |
| 164 |   --{* Wrong! *} 
 | |
| 14592 | 165 | oops | 
| 166 | ||
| 40645 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 167 | |
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 168 | subsection {* Exhaustive Testing beats Random Testing *}
 | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 169 | |
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 170 | 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: 
37929diff
changeset | 171 | where exhaustive testing beats random testing *} | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 172 | |
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 173 | lemma | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 174 | "[] ~= xs ==> hd xs = last (x # xs)" | 
| 40917 | 175 | quickcheck[random] | 
| 176 | quickcheck[exhaustive, expect = counterexample] | |
| 40645 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 177 | oops | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 178 | |
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 179 | lemma | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 180 | 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: 
37929diff
changeset | 181 | shows "drop n xs = takeWhile P xs" | 
| 40917 | 182 | quickcheck[random, iterations = 10000, quiet] | 
| 183 | quickcheck[exhaustive, expect = counterexample] | |
| 40645 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 184 | oops | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 185 | |
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 186 | lemma | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 187 | "i < length (List.transpose (List.transpose xs)) ==> xs ! i = map (%xs. xs ! i) [ys<-xs. i < length ys]" | 
| 40917 | 188 | quickcheck[random, iterations = 10000] | 
| 189 | quickcheck[exhaustive, expect = counterexample] | |
| 40645 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 190 | oops | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 191 | |
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 192 | lemma | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
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: 
37929diff
changeset | 196 | oops | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 197 | |
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 198 | lemma | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 199 | "i < n - m ==> f (lcm m i) = map f [m..<n] ! i" | 
| 40917 | 200 | quickcheck[random, iterations = 10000, finite_types = false] | 
| 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: 
37929diff
changeset | 202 | oops | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 203 | |
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 204 | lemma | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 205 | "ns ! k < length ns ==> k <= listsum ns" | 
| 40917 | 206 | quickcheck[random, iterations = 10000, finite_types = false, quiet] | 
| 207 | 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: 
37929diff
changeset | 208 | oops | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 209 | |
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 210 | lemma | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 211 | "[| ys = x # xs1; zs = xs1 @ xs |] ==> ys @ zs = x # 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: 
37929diff
changeset | 214 | oops | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 215 | |
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 216 | lemma | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 217 | "i < length xs ==> take (Suc i) xs = [] @ xs ! i # take i xs" | 
| 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: 
37929diff
changeset | 220 | oops | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 221 | |
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 222 | lemma | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 223 | "i < length xs ==> take (Suc i) xs = (xs ! i # xs) @ take i []" | 
| 40917 | 224 | quickcheck[random, iterations = 10000] | 
| 225 | quickcheck[exhaustive, expect = counterexample] | |
| 40645 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 226 | oops | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 227 | |
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 228 | lemma | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 229 | "[| sorted (rev (map length xs)); i < length xs |] ==> xs ! i = map (%ys. ys ! i) [ys<-remdups 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: 
37929diff
changeset | 232 | oops | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 233 | |
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 234 | lemma | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 235 |   "[| sorted (rev (map length xs)); i < length xs |] ==> xs ! i = map (%ys. ys ! i) [ys<-List.transpose xs. {..<i} (length ys)]"
 | 
| 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: 
37929diff
changeset | 238 | oops | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 239 | |
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 240 | lemma | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 241 | "(ys = zs) = (xs @ ys = splice xs zs)" | 
| 40917 | 242 | quickcheck[random] | 
| 243 | quickcheck[exhaustive, expect = counterexample] | |
| 40645 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 244 | oops | 
| 
03ce94672ee6
adding test cases for smallcheck and adding examples where exhaustive testing is more successful
 bulwahn parents: 
37929diff
changeset | 245 | |
| 41231 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 246 | subsection {* Examples with quantifiers *}
 | 
| 40654 | 247 | |
| 248 | text {*
 | |
| 249 | These examples show that we can handle quantifiers. | |
| 250 | *} | |
| 251 | ||
| 252 | lemma "(\<exists>x. P x) \<longrightarrow> (\<forall>x. P x)" | |
| 40917 | 253 | quickcheck[random, expect = counterexample] | 
| 254 | quickcheck[exhaustive, expect = counterexample] | |
| 40654 | 255 | oops | 
| 256 | ||
| 257 | lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>y. \<forall>x. P x y)" | |
| 40917 | 258 | quickcheck[random, expect = counterexample] | 
| 40654 | 259 | quickcheck[expect = counterexample] | 
| 260 | oops | |
| 261 | ||
| 262 | lemma "(\<exists>x. P x) \<longrightarrow> (EX! x. P x)" | |
| 40917 | 263 | quickcheck[random, expect = counterexample] | 
| 40654 | 264 | quickcheck[expect = counterexample] | 
| 265 | oops | |
| 266 | ||
| 41231 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 267 | subsection {* Examples with numerical types *}
 | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 268 | |
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 269 | text {*
 | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 270 | Quickcheck supports the common types nat, int, rat and real. | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 271 | *} | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 272 | |
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 273 | lemma | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 274 | "(x :: nat) > 0 ==> y > 0 ==> z > 0 ==> x * x + y * y \<noteq> z * z" | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 275 | quickcheck[exhaustive, size = 10, expect = counterexample] | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 276 | quickcheck[random, size = 10] | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 277 | oops | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 278 | |
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 279 | lemma | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 280 | "(x :: int) > 0 ==> y > 0 ==> z > 0 ==> x * x + y * y \<noteq> z * z" | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 281 | quickcheck[exhaustive, size = 10, expect = counterexample] | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 282 | quickcheck[random, size = 10] | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 283 | oops | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 284 | |
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 285 | lemma | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 286 | "(x :: rat) > 0 ==> y > 0 ==> z > 0 ==> x * x + y * y \<noteq> z * z" | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 287 | quickcheck[exhaustive, size = 10, expect = counterexample] | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 288 | quickcheck[random, size = 10] | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 289 | oops | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 290 | |
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 291 | lemma | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 292 | "(x :: real) > 0 ==> y > 0 ==> z > 0 ==> x * x + y * y \<noteq> z * z" | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 293 | quickcheck[exhaustive, size = 10, expect = counterexample] | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 294 | quickcheck[random, size = 10] | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 295 | oops | 
| 
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
 bulwahn parents: 
40917diff
changeset | 296 | |
| 42434 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 297 | subsection {* Examples with locales *}
 | 
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 298 | |
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 299 | locale Truth | 
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 300 | |
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 301 | context Truth | 
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 302 | begin | 
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 303 | |
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 304 | lemma "False" | 
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 305 | quickcheck[expect = no_counterexample] | 
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 306 | oops | 
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 307 | |
| 14592 | 308 | end | 
| 42434 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 309 | |
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 310 | interpretation Truth . | 
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 311 | |
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 312 | context Truth | 
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 313 | begin | 
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 314 | |
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 315 | lemma "False" | 
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 316 | quickcheck[expect = counterexample] | 
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 317 | oops | 
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 318 | |
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 319 | end | 
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 320 | |
| 
1914fd5d7c0e
adding examples for Quickcheck used within locales
 bulwahn parents: 
42159diff
changeset | 321 | end |