src/HOL/Nitpick_Examples/Pattern_Nits.thy
changeset 35284 9edc2bd6d2bd
parent 35078 6fd1052fe463
child 40341 03156257040f
--- a/src/HOL/Nitpick_Examples/Pattern_Nits.thy	Mon Feb 22 14:36:10 2010 +0100
+++ b/src/HOL/Nitpick_Examples/Pattern_Nits.thy	Mon Feb 22 19:31:00 2010 +0100
@@ -11,8 +11,8 @@
 imports Main
 begin
 
-nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60 s,
-                card = 14]
+nitpick_params [card = 10, max_potential = 0, sat_solver = MiniSat_JNI,
+                max_threads = 1, timeout = 60 s]
 
 lemma "x = (case u of () \<Rightarrow> y)"
 nitpick [expect = genuine]
@@ -132,7 +132,7 @@
 nitpick [expect = genuine]
 oops
 
-lemma "\<exists>y a b zs. x = (y # Some (a, b) # zs)"
+lemma "\<exists>y a b zs. x = y # Some (a, b) # zs"
 nitpick [expect = genuine]
 oops