src/HOL/ex/Quickcheck_Narrowing_Examples.thy
changeset 42024 51df23535105
parent 42023 1bd4b6d1c482
child 42184 1d4fae76ba5e
--- a/src/HOL/ex/Quickcheck_Narrowing_Examples.thy	Fri Mar 18 18:19:42 2011 +0100
+++ b/src/HOL/ex/Quickcheck_Narrowing_Examples.thy	Fri Mar 18 18:19:42 2011 +0100
@@ -42,13 +42,32 @@
 declare [[quickcheck_finite_functions]]
 
 lemma "map f xs = map g xs"
-  quickcheck[tester = narrowing, finite_types = true]
+  quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
 oops
 
 lemma "map f xs = map f ys ==> xs = ys"
-  quickcheck[tester = narrowing, finite_types = true]
+  quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
+oops
+
+lemma
+  "list_all2 P (rev xs) (rev ys) = list_all2 P xs (rev ys)"
+  quickcheck[tester = narrowing, expect = counterexample]
+oops
+
+lemma "map f xs = F f xs"
+  quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
 oops
 
+lemma "map f xs = F f xs"
+  quickcheck[tester = narrowing, finite_types = false, expect = counterexample]
+oops
+
+(* requires some work...
+lemma "R O S = S O R"
+  quickcheck[tester = narrowing, size = 2]
+oops
+*)
+
 subsection {* AVL Trees *}
 
 datatype 'a tree = ET |  MKT 'a "'a tree" "'a tree" nat