diff -r 4b198af5beb5 -r c57dba973391 src/HOL/Nitpick.thy --- a/src/HOL/Nitpick.thy Sat Feb 13 11:56:06 2010 +0100 +++ b/src/HOL/Nitpick.thy Sat Feb 13 15:04:09 2010 +0100 @@ -37,7 +37,6 @@ and bisim_iterator_max :: bisim_iterator and Quot :: "'a \ 'b" and quot_normal :: "'a \ 'a" - and Silly :: "'a \ 'b" and Tha :: "('a \ bool) \ 'a" datatype ('a, 'b) pair_box = PairBox 'a 'b @@ -45,7 +44,6 @@ typedecl unsigned_bit typedecl signed_bit -typedecl \ datatype 'a word = Word "('a set)" @@ -254,12 +252,12 @@ setup {* Nitpick_Isar.setup *} hide (open) const unknown is_unknown undefined_fast_The undefined_fast_Eps bisim - bisim_iterator_max Quot quot_normal Silly Tha PairBox FunBox Word refl' wf' + bisim_iterator_max Quot quot_normal Tha PairBox FunBox Word refl' wf' wf_wfrec wf_wfrec' wfrec' card' setsum' fold_graph' nat_gcd nat_lcm int_gcd int_lcm Frac Abs_Frac Rep_Frac zero_frac one_frac num denom norm_frac frac plus_frac times_frac uminus_frac number_of_frac inverse_frac less_eq_frac of_frac -hide (open) type bisim_iterator pair_box fun_box unsigned_bit signed_bit \ word +hide (open) type bisim_iterator pair_box fun_box unsigned_bit signed_bit word hide (open) fact If_def Ex1_def rtrancl_def rtranclp_def tranclp_def refl'_def wf'_def wf_wfrec'_def wfrec'_def card'_def setsum'_def fold_graph'_def The_psimp Eps_psimp unit_case_def nat_case_def list_size_simp nat_gcd_def