diff -r fbf77fdf9ae4 -r 359bec38a4ee src/HOL/Nitpick_Examples/Manual_Nits.thy --- a/src/HOL/Nitpick_Examples/Manual_Nits.thy Wed May 30 16:59:20 2012 +0200 +++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Wed May 30 18:09:23 2012 +0200 @@ -64,8 +64,10 @@ subsection {* 2.5. Natural Numbers and Integers *} lemma "\i \ j; n \ (m\int)\ \ i * n + j * m \ i * m + j * n" +(* FIXME nitpick [expect = genuine] nitpick [binary_ints, bits = 16, expect = genuine] +*) oops lemma "\n. Suc n \ n \ P" @@ -141,11 +143,15 @@ Ycoord :: int lemma "Xcoord (p\point) = Xcoord (q\point)" +(* FIXME: Invalid intermediate term nitpick [show_datatypes, expect = genuine] +*) oops lemma "4 * x + 3 * (y\real) \ 1 / 2" +(* FIXME: Invalid intermediate term nitpick [show_datatypes, expect = genuine] +*) oops subsection {* 2.8. Inductive and Coinductive Predicates *}