src/HOL/Nitpick.thy
changeset 35671 ed2c3830d881
parent 35665 ff2bf50505ab
child 35699 9ed327529a44
--- a/src/HOL/Nitpick.thy	Tue Mar 09 09:25:23 2010 +0100
+++ b/src/HOL/Nitpick.thy	Tue Mar 09 14:18:21 2010 +0100
@@ -36,7 +36,8 @@
            and bisim :: "bisim_iterator \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
            and bisim_iterator_max :: bisim_iterator
            and Quot :: "'a \<Rightarrow> 'b"
-           and Tha :: "('a \<Rightarrow> bool) \<Rightarrow> 'a"
+           and safe_The :: "('a \<Rightarrow> bool) \<Rightarrow> 'a"
+           and safe_Eps :: "('a \<Rightarrow> bool) \<Rightarrow> 'a"
 
 datatype ('a, 'b) fin_fun = FinFun "('a \<Rightarrow> 'b)"
 datatype ('a, 'b) fun_box = FunBox "('a \<Rightarrow> 'b)"
@@ -237,10 +238,11 @@
 setup {* Nitpick_Isar.setup *}
 
 hide (open) const unknown is_unknown undefined_fast_The undefined_fast_Eps bisim 
-    bisim_iterator_max Quot Tha FinFun FunBox PairBox 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
+    bisim_iterator_max Quot safe_The safe_Eps FinFun FunBox PairBox 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 fin_fun fun_box pair_box unsigned_bit signed_bit
     word
 hide (open) fact If_def Ex1_def rtrancl_def rtranclp_def tranclp_def refl'_def