| author | wenzelm | 
| Sun, 05 Sep 2010 19:47:40 +0200 | |
| changeset 39133 | 70d3915c92f0 | 
| parent 37929 | 22e0797857e6 | 
| child 40645 | 03ce94672ee6 | 
| permissions | -rw-r--r-- | 
| 14592 | 1 | (* Title: HOL/ex/Quickcheck_Examples.thy | 
| 2 | Author: Stefan Berghofer | |
| 3 | Copyright 2004 TU Muenchen | |
| 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 | |
| 14 | formulae under an assignment of free variables to random values. | |
| 15 | In contrast to 'refute', it can deal with inductive datatypes, | |
| 16 | but cannot handle quantifiers. | |
| 17 | *} | |
| 18 | ||
| 19 | subsection {* Lists *}
 | |
| 20 | ||
| 21 | theorem "map g (map f xs) = map (g o f) xs" | |
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 22 | quickcheck[expect = no_counterexample] | 
| 14592 | 23 | oops | 
| 24 | ||
| 25 | theorem "map g (map f xs) = map (f o g) xs" | |
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 26 | quickcheck[expect = counterexample] | 
| 14592 | 27 | oops | 
| 28 | ||
| 29 | theorem "rev (xs @ ys) = rev ys @ rev xs" | |
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 30 | quickcheck[expect = no_counterexample] | 
| 14592 | 31 | oops | 
| 32 | ||
| 33 | theorem "rev (xs @ ys) = rev xs @ rev ys" | |
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 34 | quickcheck[expect = counterexample] | 
| 14592 | 35 | oops | 
| 36 | ||
| 37 | theorem "rev (rev xs) = xs" | |
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 38 | quickcheck[expect = no_counterexample] | 
| 14592 | 39 | oops | 
| 40 | ||
| 41 | theorem "rev xs = xs" | |
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 42 | quickcheck[expect = counterexample] | 
| 14592 | 43 | oops | 
| 44 | ||
| 25891 | 45 | text {* An example involving functions inside other data structures *}
 | 
| 46 | ||
| 28314 | 47 | primrec app :: "('a \<Rightarrow> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a" where
 | 
| 25891 | 48 | "app [] x = x" | 
| 28314 | 49 | | "app (f # fs) x = app fs (f x)" | 
| 25891 | 50 | |
| 51 | lemma "app (fs @ gs) x = app gs (app fs x)" | |
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 52 | quickcheck[expect = no_counterexample] | 
| 25891 | 53 | by (induct fs arbitrary: x) simp_all | 
| 54 | ||
| 55 | lemma "app (fs @ gs) x = app fs (app gs x)" | |
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 56 | quickcheck[expect = counterexample] | 
| 25891 | 57 | oops | 
| 58 | ||
| 28314 | 59 | primrec occurs :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where | 
| 14592 | 60 | "occurs a [] = 0" | 
| 28314 | 61 | | "occurs a (x#xs) = (if (x=a) then Suc(occurs a xs) else occurs a xs)" | 
| 14592 | 62 | |
| 28314 | 63 | primrec del1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where | 
| 14592 | 64 | "del1 a [] = []" | 
| 28314 | 65 | | "del1 a (x#xs) = (if (x=a) then xs else (x#del1 a xs))" | 
| 14592 | 66 | |
| 25891 | 67 | text {* A lemma, you'd think to be true from our experience with delAll *}
 | 
| 14592 | 68 | lemma "Suc (occurs a (del1 a xs)) = occurs a xs" | 
| 69 |   -- {* Wrong. Precondition needed.*}
 | |
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 70 | quickcheck[expect = counterexample] | 
| 14592 | 71 | oops | 
| 72 | ||
| 73 | lemma "xs ~= [] \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs" | |
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 74 | quickcheck[expect = counterexample] | 
| 14592 | 75 |     -- {* Also wrong.*}
 | 
| 76 | oops | |
| 77 | ||
| 78 | lemma "0 < occurs a xs \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs" | |
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 79 | quickcheck[expect = no_counterexample] | 
| 28314 | 80 | by (induct xs) auto | 
| 14592 | 81 | |
| 28314 | 82 | primrec replace :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where | 
| 14592 | 83 | "replace a b [] = []" | 
| 28314 | 84 | | "replace a b (x#xs) = (if (x=a) then (b#(replace a b xs)) | 
| 14592 | 85 | else (x#(replace a b xs)))" | 
| 86 | ||
| 87 | lemma "occurs a xs = occurs b (replace a b xs)" | |
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 88 | quickcheck[expect = counterexample] | 
| 14592 | 89 |   -- {* Wrong. Precondition needed.*}
 | 
| 90 | oops | |
| 91 | ||
| 92 | lemma "occurs b xs = 0 \<or> a=b \<longrightarrow> occurs a xs = occurs b (replace a b xs)" | |
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 93 | quickcheck[expect = no_counterexample] | 
| 28314 | 94 | by (induct xs) simp_all | 
| 14592 | 95 | |
| 96 | ||
| 97 | subsection {* Trees *}
 | |
| 98 | ||
| 99 | datatype 'a tree = Twig | Leaf 'a | Branch "'a tree" "'a tree" | |
| 100 | ||
| 28314 | 101 | primrec leaves :: "'a tree \<Rightarrow> 'a list" where | 
| 14592 | 102 | "leaves Twig = []" | 
| 28314 | 103 | | "leaves (Leaf a) = [a]" | 
| 104 | | "leaves (Branch l r) = (leaves l) @ (leaves r)" | |
| 14592 | 105 | |
| 28314 | 106 | primrec plant :: "'a list \<Rightarrow> 'a tree" where | 
| 14592 | 107 | "plant [] = Twig " | 
| 28314 | 108 | | "plant (x#xs) = Branch (Leaf x) (plant xs)" | 
| 14592 | 109 | |
| 28314 | 110 | primrec mirror :: "'a tree \<Rightarrow> 'a tree" where | 
| 14592 | 111 | "mirror (Twig) = Twig " | 
| 28314 | 112 | | "mirror (Leaf a) = Leaf a " | 
| 113 | | "mirror (Branch l r) = Branch (mirror r) (mirror l)" | |
| 14592 | 114 | |
| 115 | theorem "plant (rev (leaves xt)) = mirror xt" | |
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 116 | quickcheck[expect = counterexample] | 
| 14592 | 117 |     --{* Wrong! *} 
 | 
| 118 | oops | |
| 119 | ||
| 120 | theorem "plant((leaves xt) @ (leaves yt)) = Branch xt yt" | |
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 121 | quickcheck[expect = counterexample] | 
| 14592 | 122 |     --{* Wrong! *} 
 | 
| 123 | oops | |
| 124 | ||
| 125 | datatype 'a ntree = Tip "'a" | Node "'a" "'a ntree" "'a ntree" | |
| 126 | ||
| 28314 | 127 | primrec inOrder :: "'a ntree \<Rightarrow> 'a list" where | 
| 14592 | 128 | "inOrder (Tip a)= [a]" | 
| 28314 | 129 | | "inOrder (Node f x y) = (inOrder x)@[f]@(inOrder y)" | 
| 14592 | 130 | |
| 28314 | 131 | primrec root :: "'a ntree \<Rightarrow> 'a" where | 
| 14592 | 132 | "root (Tip a) = a" | 
| 28314 | 133 | | "root (Node f x y) = f" | 
| 14592 | 134 | |
| 28314 | 135 | theorem "hd (inOrder xt) = root xt" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37914diff
changeset | 136 | quickcheck[expect = counterexample] | 
| 14592 | 137 |     --{* Wrong! *} 
 | 
| 138 | oops | |
| 139 | ||
| 140 | end |