--- a/src/HOL/ex/Quickcheck_Examples.thy Mon Sep 22 13:56:01 2008 +0200
+++ b/src/HOL/ex/Quickcheck_Examples.thy Mon Sep 22 13:56:03 2008 +0200
@@ -6,7 +6,9 @@
header {* Examples for the 'quickcheck' command *}
-theory Quickcheck_Examples imports Main begin
+theory Quickcheck_Examples
+imports Main
+begin
text {*
The 'quickcheck' command allows to find counterexamples by evaluating
@@ -18,144 +20,121 @@
subsection {* Lists *}
theorem "map g (map f xs) = map (g o f) xs"
- quickcheck
+ quickcheck []
oops
theorem "map g (map f xs) = map (f o g) xs"
- quickcheck
+ quickcheck []
oops
theorem "rev (xs @ ys) = rev ys @ rev xs"
- quickcheck
+ quickcheck []
oops
theorem "rev (xs @ ys) = rev xs @ rev ys"
- quickcheck
+ quickcheck []
oops
theorem "rev (rev xs) = xs"
- quickcheck
+ quickcheck []
oops
theorem "rev xs = xs"
- quickcheck
+ quickcheck []
oops
text {* An example involving functions inside other data structures *}
-consts app :: "('a \<Rightarrow> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a"
-
-primrec
+primrec app :: "('a \<Rightarrow> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a" where
"app [] x = x"
- "app (f # fs) x = app fs (f x)"
+ | "app (f # fs) x = app fs (f x)"
lemma "app (fs @ gs) x = app gs (app fs x)"
- quickcheck
+ quickcheck []
by (induct fs arbitrary: x) simp_all
lemma "app (fs @ gs) x = app fs (app gs x)"
- quickcheck
+ quickcheck []
oops
-consts
- occurs :: "'a \<Rightarrow> 'a list \<Rightarrow> nat"
-primrec
+primrec occurs :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where
"occurs a [] = 0"
- "occurs a (x#xs) = (if (x=a) then Suc(occurs a xs) else occurs a xs)"
+ | "occurs a (x#xs) = (if (x=a) then Suc(occurs a xs) else occurs a xs)"
-consts
- del1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list"
-primrec
+primrec del1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
"del1 a [] = []"
- "del1 a (x#xs) = (if (x=a) then xs else (x#del1 a xs))"
+ | "del1 a (x#xs) = (if (x=a) then xs else (x#del1 a xs))"
text {* A lemma, you'd think to be true from our experience with delAll *}
lemma "Suc (occurs a (del1 a xs)) = occurs a xs"
-- {* Wrong. Precondition needed.*}
- quickcheck
+ quickcheck []
oops
lemma "xs ~= [] \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
- quickcheck
+ quickcheck []
-- {* Also wrong.*}
oops
lemma "0 < occurs a xs \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
- quickcheck
- apply (induct_tac xs)
- apply auto
- -- {* Correct! *}
- done
+ quickcheck []
+ by (induct xs) auto
-consts
- replace :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list"
-primrec
+primrec replace :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
"replace a b [] = []"
- "replace a b (x#xs) = (if (x=a) then (b#(replace a b xs))
+ | "replace a b (x#xs) = (if (x=a) then (b#(replace a b xs))
else (x#(replace a b xs)))"
lemma "occurs a xs = occurs b (replace a b xs)"
- quickcheck
+ quickcheck []
-- {* Wrong. Precondition needed.*}
oops
lemma "occurs b xs = 0 \<or> a=b \<longrightarrow> occurs a xs = occurs b (replace a b xs)"
- quickcheck
- apply (induct_tac xs)
- apply auto
- done
+ quickcheck []
+ by (induct xs) simp_all
subsection {* Trees *}
datatype 'a tree = Twig | Leaf 'a | Branch "'a tree" "'a tree"
-consts
- leaves :: "'a tree \<Rightarrow> 'a list"
-primrec
+primrec leaves :: "'a tree \<Rightarrow> 'a list" where
"leaves Twig = []"
- "leaves (Leaf a) = [a]"
- "leaves (Branch l r) = (leaves l) @ (leaves r)"
+ | "leaves (Leaf a) = [a]"
+ | "leaves (Branch l r) = (leaves l) @ (leaves r)"
-consts
- plant :: "'a list \<Rightarrow> 'a tree"
-primrec
+primrec plant :: "'a list \<Rightarrow> 'a tree" where
"plant [] = Twig "
- "plant (x#xs) = Branch (Leaf x) (plant xs)"
+ | "plant (x#xs) = Branch (Leaf x) (plant xs)"
-consts
- mirror :: "'a tree \<Rightarrow> 'a tree"
-primrec
+primrec mirror :: "'a tree \<Rightarrow> 'a tree" where
"mirror (Twig) = Twig "
- "mirror (Leaf a) = Leaf a "
- "mirror (Branch l r) = Branch (mirror r) (mirror l)"
+ | "mirror (Leaf a) = Leaf a "
+ | "mirror (Branch l r) = Branch (mirror r) (mirror l)"
theorem "plant (rev (leaves xt)) = mirror xt"
- quickcheck
+ quickcheck []
--{* Wrong! *}
oops
theorem "plant((leaves xt) @ (leaves yt)) = Branch xt yt"
- quickcheck
+ quickcheck []
--{* Wrong! *}
oops
datatype 'a ntree = Tip "'a" | Node "'a" "'a ntree" "'a ntree"
-consts
- inOrder :: "'a ntree \<Rightarrow> 'a list"
-primrec
+primrec inOrder :: "'a ntree \<Rightarrow> 'a list" where
"inOrder (Tip a)= [a]"
- "inOrder (Node f x y) = (inOrder x)@[f]@(inOrder y)"
+ | "inOrder (Node f x y) = (inOrder x)@[f]@(inOrder y)"
-consts
- root :: "'a ntree \<Rightarrow> 'a"
-primrec
+primrec root :: "'a ntree \<Rightarrow> 'a" where
"root (Tip a) = a"
- "root (Node f x y) = f"
+ | "root (Node f x y) = f"
-theorem "hd(inOrder xt) = root xt"
- quickcheck
+theorem "hd (inOrder xt) = root xt"
+ quickcheck []
--{* Wrong! *}
oops