src/HOL/Library/Efficient_Nat.thy
changeset 29932 a2594b5c945a
parent 29815 9e94b7078fa5
child 29937 ffed4bd4bfad
--- a/src/HOL/Library/Efficient_Nat.thy	Mon Feb 16 13:38:08 2009 +0100
+++ b/src/HOL/Library/Efficient_Nat.thy	Mon Feb 16 13:38:09 2009 +0100
@@ -107,10 +107,10 @@
 
 lemma Suc_if_eq: "(\<And>n. f (Suc n) = h n) \<Longrightarrow> f 0 = g \<Longrightarrow>
   f n = (if n = 0 then g else h (n - 1))"
-  by (case_tac n) simp_all
+  by (cases n) simp_all
 
 lemma Suc_clause: "(\<And>n. P n (Suc n)) \<Longrightarrow> n \<noteq> 0 \<Longrightarrow> P (n - 1) n"
-  by (case_tac n) simp_all
+  by (cases n) simp_all
 
 text {*
   The rules above are built into a preprocessor that is plugged into
@@ -132,7 +132,7 @@
       (fst (Thm.dest_comb (snd (Thm.dest_comb (cprop_of th))))));
     fun rhs_of th = snd (Thm.dest_comb (snd (Thm.dest_comb (cprop_of th))));
     fun find_vars ct = (case term_of ct of
-        (Const ("Suc", _) $ Var _) => [(cv, snd (Thm.dest_comb ct))]
+        (Const (@{const_name Suc}, _) $ Var _) => [(cv, snd (Thm.dest_comb ct))]
       | _ $ _ =>
         let val (ct1, ct2) = Thm.dest_comb ct
         in 
@@ -236,7 +236,6 @@
   Codegen.add_preprocessor eqn_suc_preproc
   #> Codegen.add_preprocessor clause_suc_preproc
   #> Code.add_functrans ("eqn_Suc", lift eqn_suc_preproc)
-  #> Code.add_functrans ("clause_Suc", lift clause_suc_preproc)
 
 end;
 *}