src/HOL/ex/Quickcheck_Examples.thy
changeset 28336 a8edf4c69a79
parent 28314 053419cefd3c
child 37914 49b908e43d61
--- a/src/HOL/ex/Quickcheck_Examples.thy	Tue Sep 23 18:11:42 2008 +0200
+++ b/src/HOL/ex/Quickcheck_Examples.thy	Tue Sep 23 18:11:43 2008 +0200
@@ -20,27 +20,27 @@
 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 *}
@@ -50,11 +50,11 @@
   | "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
 
 primrec occurs :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where
@@ -68,16 +68,16 @@
 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 []
+  quickcheck
   by (induct xs) auto
 
 primrec replace :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
@@ -86,12 +86,12 @@
                             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 []
+  quickcheck
   by (induct xs) simp_all
 
 
@@ -114,12 +114,12 @@
   | "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
 
@@ -134,7 +134,7 @@
   | "root (Node f x y) = f"
 
 theorem "hd (inOrder xt) = root xt"
-  quickcheck []
+  quickcheck
     --{* Wrong! *} 
   oops