src/HOL/Nitpick.thy
changeset 55415 05f5fdb8d093
parent 55414 eab03e9cee8a
child 55539 0819931d652d
--- a/src/HOL/Nitpick.thy	Wed Feb 12 08:35:57 2014 +0100
+++ b/src/HOL/Nitpick.thy	Wed Feb 12 08:35:57 2014 +0100
@@ -104,8 +104,8 @@
 
 declare unit.cases [nitpick_simp del]
 
-lemma nat_case_unfold [nitpick_unfold]:
-"nat_case x f n \<equiv> if n = 0 then x else f (n - 1)"
+lemma case_nat_unfold [nitpick_unfold]:
+"case_nat x f n \<equiv> if n = 0 then x else f (n - 1)"
 apply (rule eq_reflection)
 by (cases n) auto
 
@@ -241,7 +241,7 @@
 hide_type (open) bisim_iterator fun_box pair_box unsigned_bit signed_bit word
 hide_fact (open) Ex1_unfold rtrancl_unfold rtranclp_unfold tranclp_unfold
     prod_def refl'_def wf'_def card'_def setsum'_def
-    fold_graph'_def The_psimp Eps_psimp case_unit_unfold nat_case_unfold
+    fold_graph'_def The_psimp Eps_psimp case_unit_unfold case_nat_unfold
     list_size_simp nat_gcd_def nat_lcm_def int_gcd_def int_lcm_def Frac_def
     zero_frac_def one_frac_def num_def denom_def norm_frac_def frac_def
     plus_frac_def times_frac_def uminus_frac_def number_of_frac_def