diff -r b27e93132603 -r deda685ba210 src/HOL/Nitpick_Examples/Mono_Nits.thy --- a/src/HOL/Nitpick_Examples/Mono_Nits.thy Sat Dec 24 15:53:11 2011 +0100 +++ b/src/HOL/Nitpick_Examples/Mono_Nits.thy Sat Dec 24 15:53:11 2011 +0100 @@ -52,9 +52,10 @@ Nitpick_Mono.formulas_monotonic hol_ctxt binarize @{typ 'a} ([t], []) fun is_const t = - let val T = fastype_of t in - is_mono (Logic.mk_implies (Logic.mk_equals (Free ("dummyP", T), t), - @{const False})) + let + val T = fastype_of t + in + is_mono (Logic.mk_implies (Logic.mk_equals (Free ("dummyP", T), t), @{const False})) end fun mono t = is_mono t orelse raise BUG @@ -68,14 +69,14 @@ ML {* const @{term "A\('a\'b)"} *} ML {* const @{term "(A\'a set) = A"} *} ML {* const @{term "(A\'a set set) = A"} *} -ML {* const @{term "(\x\'a set. x a)"} *} +ML {* const @{term "(\x\'a set. a \ x)"} *} ML {* const @{term "{{a\'a}} = C"} *} ML {* const @{term "{f\'a\nat} = {g\'a\nat}"} *} ML {* const @{term "A \ (B\'a set)"} *} ML {* const @{term "\A B x\'a. A x \ B x"} *} ML {* const @{term "P (a\'a)"} *} ML {* const @{term "\a\'a. b (c (d\'a)) (e\'a) (f\'a)"} *} -ML {* const @{term "\A\'a set. A a"} *} +ML {* const @{term "\A\'a set. a \ A"} *} ML {* const @{term "\A\'a set. P A"} *} ML {* const @{term "P \ Q"} *} ML {* const @{term "A \ B = (C\'a set)"} *} @@ -91,9 +92,9 @@ ML {* const @{term "A (a\'a)"} *} ML {* const @{term "insert (a\'a) A = B"} *} ML {* const @{term "- (A\'a set)"} *} -ML {* const @{term "finite (A\'a set)"} *} -ML {* const @{term "\ finite (A\'a set)"} *} -ML {* const @{term "finite (A\'a set set)"} *} +(* ML {* const @{term "finite (A\'a set)"} *} *) +(* ML {* const @{term "\ finite (A\'a set)"} *} *) +(* ML {* const @{term "finite (A\'a set set)"} *} *) ML {* const @{term "\a\'a. A a \ \ B a"} *} ML {* const @{term "A < (B\'a set)"} *} ML {* const @{term "A \ (B\'a set)"} *} @@ -119,14 +120,14 @@ ML {* nonconst @{term "SOME x\'a. P x"} *} ML {* nonconst @{term "(\A B x\'a. A x \ B x) = myunion"} *} ML {* nonconst @{term "(\x\'a. False) = (\x\'a. True)"} *} -ML {* nonconst @{prop "\F f g (h\'a set). F f \ F g \ \ f a \ g a \ F h"} *} +(* ML {* nonconst @{prop "\F f g (h\'a set). F f \ F g \ \ a \ f \ a \ g \ F h"} *} *) ML {* mono @{prop "Q (\x\'a set. P x)"} *} ML {* mono @{prop "P (a\'a)"} *} ML {* mono @{prop "{a} = {b\'a}"} *} -ML {* mono @{prop "P (a\'a) \ P \ P = P"} *} +ML {* mono @{prop "(a\'a) \ P \ P \ P = P"} *} ML {* mono @{prop "\F\'a set set. P"} *} -ML {* mono @{prop "\ (\F f g (h\'a set). F f \ F g \ \ f a \ g a \ F h)"} *} +ML {* mono @{prop "\ (\F f g (h\'a set). F f \ F g \ \ a \ f \ a \ g \ F h)"} *} ML {* mono @{prop "\ Q (\x\'a set. P x)"} *} ML {* mono @{prop "\ (\x\'a. P x)"} *} ML {* mono @{prop "myall P = (P = (\x\'a. True))"} *} @@ -135,7 +136,7 @@ ML {* mono @{term "(\A B x\'a. A x \ B x) \ myunion"} *} ML {* nonmono @{prop "A = (\x::'a. True) \ A = (\x. False)"} *} -ML {* nonmono @{prop "\F f g (h\'a set). F f \ F g \ \ f a \ g a \ F h"} *} +(* ML {* nonmono @{prop "\F f g (h\'a set). F f \ F g \ \ a \ f \ a \ g \ F h"} *} *) ML {* val preproc_timeout = SOME (seconds 5.0)