diff -r 0681e46b4022 -r e482320bcbfe src/HOL/Nitpick_Examples/Manual_Nits.thy --- a/src/HOL/Nitpick_Examples/Manual_Nits.thy Mon Jun 21 11:15:21 2010 +0200 +++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Mon Jun 21 11:16:00 2010 +0200 @@ -7,13 +7,17 @@ header {* Examples from the Nitpick Manual *} +(* The "expect" arguments to Nitpick in this theory and the other example + theories are there so that the example can also serve as a regression test + suite. *) + theory Manual_Nits imports Main Quotient_Product RealDef begin chapter {* 3. First Steps *} -nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1] +nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60 s] subsection {* 3.1. Propositional Logic *} @@ -379,7 +383,7 @@ theorem S\<^isub>3_sound: "w \ S\<^isub>3 \ length [x \ w. x = a] = length [x \ w. x = b]" -nitpick [card = 1\6, expect = none] +nitpick [card = 1\5, expect = none] sorry theorem S\<^isub>3_complete: @@ -398,19 +402,19 @@ theorem S\<^isub>4_sound: "w \ S\<^isub>4 \ length [x \ w. x = a] = length [x \ w. x = b]" -nitpick [card = 1\6, expect = none] +nitpick [card = 1\5, expect = none] sorry theorem S\<^isub>4_complete: "length [x \ w. x = a] = length [x \ w. x = b] \ w \ S\<^isub>4" -nitpick [card = 1\6, expect = none] +nitpick [card = 1\5, expect = none] sorry theorem S\<^isub>4_A\<^isub>4_B\<^isub>4_sound_and_complete: "w \ S\<^isub>4 \ length [x \ w. x = a] = length [x \ w. x = b]" "w \ A\<^isub>4 \ length [x \ w. x = a] = length [x \ w. x = b] + 1" "w \ B\<^isub>4 \ length [x \ w. x = b] = length [x \ w. x = a] + 1" -nitpick [card = 1\6, expect = none] +nitpick [card = 1\5, expect = none] sorry subsection {* 4.2. AA Trees *}