merged
authorAndreas Lochbihler
Wed, 12 Feb 2014 10:59:25 +0100
changeset 55428 0ab52bf7b5e6
parent 55427 ff54d22fe357 (current diff)
parent 55425 7a3e78ee813b (diff)
child 55440 721b4561007a
merged
src/HOL/Code_Numeral.thy
src/HOL/String.thy
--- a/NEWS	Wed Feb 12 09:06:04 2014 +0100
+++ b/NEWS	Wed Feb 12 10:59:25 2014 +0100
@@ -122,6 +122,11 @@
     BNF/BNF.thy
     BNF/Equiv_Relations_More.thy
 
+* Old datatype package:
+  * Generated constants "xxx_case" and "xxx_rec" have been renamed "case_xxx"
+    and "rec_xxx".
+    INCOMPATIBILITY.
+
 * New theory:
     Cardinals/Ordinal_Arithmetic.thy
 
--- a/src/Doc/Codegen/Evaluation.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/Doc/Codegen/Evaluation.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -296,7 +296,7 @@
 
 code_reflect %quote Sum_Type
   datatypes sum = Inl | Inr
-  functions "Sum_Type.Projl" "Sum_Type.Projr"
+  functions "Sum_Type.sum.projl" "Sum_Type.sum.projr"
 
 text {*
   \noindent @{command_def code_reflect} takes a structure name and
--- a/src/Doc/Codegen/Refinement.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/Doc/Codegen/Refinement.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -148,15 +148,15 @@
   \noindent It is good style, although no absolute requirement, to
   provide code equations for the original artefacts of the implemented
   type, if possible; in our case, these are the datatype constructor
-  @{const Queue} and the case combinator @{const queue_case}:
+  @{const Queue} and the case combinator @{const case_queue}:
 *}
 
 lemma %quote Queue_AQueue [code]:
   "Queue = AQueue []"
   by (simp add: AQueue_def fun_eq_iff)
 
-lemma %quote queue_case_AQueue [code]:
-  "queue_case f (AQueue xs ys) = f (ys @ rev xs)"
+lemma %quote case_queue_AQueue [code]:
+  "case_queue f (AQueue xs ys) = f (ys @ rev xs)"
   by (simp add: AQueue_def)
 
 text {*
@@ -164,7 +164,7 @@
 *}
 
 text %quotetypewriter {*
-  @{code_stmts empty enqueue dequeue Queue queue_case (SML)}
+  @{code_stmts empty enqueue dequeue Queue case_queue (SML)}
 *}
 
 text {*
@@ -274,4 +274,3 @@
 *}
 
 end
-
--- a/src/Doc/Datatypes/Datatypes.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/Doc/Datatypes/Datatypes.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -470,7 +470,7 @@
   @@{command datatype_new} target? @{syntax dt_options}? \<newline>
     (@{syntax dt_name} '=' (@{syntax ctor} + '|') + @'and')
   ;
-  @{syntax_def dt_options}: '(' (('no_discs_sels' | 'no_code' | 'rep_compat') + ',') ')'
+  @{syntax_def dt_options}: '(' (('no_discs_sels' | 'no_code') + ',') ')'
 \<close>}
 
 \medskip
@@ -492,12 +492,6 @@
 \item
 The @{text "no_code"} option indicates that the datatype should not be
 registered for code generation.
-
-\item
-The @{text "rep_compat"} option indicates that the generated names should
-contain optional (and normally not displayed) ``@{text "new."}'' components to
-prevent clashes with a later call to \keyw{rep\_datatype}. See
-Section~\ref{ssec:datatype-compatibility-issues} for details.
 \end{itemize}
 
 The left-hand sides of the datatype equations specify the name of the type to
@@ -771,7 +765,7 @@
 @{thm list.case(1)[no_vars]} \\
 @{thm list.case(2)[no_vars]}
 
-\item[@{text "t."}\hthm{case\_cong}\rm:] ~ \\
+\item[@{text "t."}\hthm{case\_cong} @{text "[fundef_cong]"}\rm:] ~ \\
 @{thm list.case_cong[no_vars]}
 
 \item[@{text "t."}\hthm{weak\_case\_cong} @{text "[cong]"}\rm:] ~ \\
@@ -2563,7 +2557,7 @@
 %    old \keyw{datatype}
 %
 %  * @{command wrap_free_constructors}
-%    * @{text "no_discs_sels"}, @{text "no_code"}, @{text "rep_compat"}
+%    * @{text "no_discs_sels"}, @{text "no_code"}
 %    * hack to have both co and nonco view via locale (cf. ext nats)
 %  * code generator
 %     * eq, refl, simps
@@ -2601,10 +2595,13 @@
 
 \medskip
 
-% options: no_discs_sels no_code rep_compat
+% options: no_discs_sels no_code
 
 \noindent
 Section~\ref{ssec:datatype-generated-theorems} lists the generated theorems.
+For technical reasons, the @{text "[fundef_cong]"} attribute is not set on the
+generated @{text case_cong} theorem. It can be added manually using
+\keyw{declare}.
 *}
 
 
--- a/src/Doc/LaTeXsugar/Sugar.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/Doc/LaTeXsugar/Sugar.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -166,9 +166,9 @@
 \end{quote}
 To support the ``\_''-notation for irrelevant variables
 the constant \texttt{DUMMY} has been introduced:
-@{thm fst_conv[where b = DUMMY]} is produced by
+@{thm fst_conv[of _ DUMMY]} is produced by
 \begin{quote}
-\verb!@!\verb!{thm fst_conv[where b = DUMMY]}!
+\verb!@!\verb!{thm fst_conv[of _ DUMMY]}!
 \end{quote}
 Variables that are bound by quantifiers or lambdas cannot be renamed
 like this. Instead, the attribute \texttt{rename\_abs} does the
--- a/src/Doc/Logics/document/HOL.tex	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/Doc/Logics/document/HOL.tex	Wed Feb 12 10:59:25 2014 +0100
@@ -1115,7 +1115,7 @@
   \it symbol    & \it meta-type &           & \it description \\ 
   \cdx{Inl}     & $\alpha \To \alpha+\beta$    & & first injection\\
   \cdx{Inr}     & $\beta \To \alpha+\beta$     & & second injection\\
-  \cdx{sum_case} & $[\alpha\To\gamma, \beta\To\gamma, \alpha+\beta] \To\gamma$
+  \cdx{case_sum} & $[\alpha\To\gamma, \beta\To\gamma, \alpha+\beta] \To\gamma$
         & & conditional
 \end{constants}
 \begin{ttbox}\makeatletter
@@ -1126,11 +1126,11 @@
 
 \tdx{sumE}           [| !!x. P(Inl x);  !!y. P(Inr y) |] ==> P s
 
-\tdx{sum_case_Inl}   sum_case f g (Inl x) = f x
-\tdx{sum_case_Inr}   sum_case f g (Inr x) = g x
+\tdx{case_sum_Inl}   case_sum f g (Inl x) = f x
+\tdx{case_sum_Inr}   case_sum f g (Inr x) = g x
 
-\tdx{surjective_sum} sum_case (\%x. f(Inl x)) (\%y. f(Inr y)) s = f s
-\tdx{sum.split_case} R(sum_case f g s) = ((! x. s = Inl(x) --> R(f(x))) &
+\tdx{surjective_sum} case_sum (\%x. f(Inl x)) (\%y. f(Inr y)) s = f s
+\tdx{sum.split_case} R(case_sum f g s) = ((! x. s = Inl(x) --> R(f(x))) &
                                      (! y. s = Inr(y) --> R(g(y))))
 \end{ttbox}
 \caption{Type $\alpha+\beta$}\label{hol-sum}
@@ -1231,7 +1231,7 @@
 Note that Isabelle insists on precisely this format; you may not even change
 the order of the two cases.
 Both \texttt{primrec} and \texttt{case} are realized by a recursion operator
-\cdx{nat_rec}, which is available because \textit{nat} is represented as
+\cdx{rec_nat}, which is available because \textit{nat} is represented as
 a datatype.
 
 %The predecessor relation, \cdx{pred_nat}, is shown to be well-founded.
@@ -1435,7 +1435,7 @@
 case $e$ of [] => $a$  |  \(x\)\#\(xs\) => b
 \end{center}
 is defined by translation.  For details see~{\S}\ref{sec:HOL:datatype}. There
-is also a case splitting rule \tdx{split_list_case}
+is also a case splitting rule \tdx{list.split}
 \[
 \begin{array}{l}
 P(\mathtt{case}~e~\mathtt{of}~\texttt{[] =>}~a ~\texttt{|}~
--- a/src/Doc/ProgProve/Isar.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/Doc/ProgProve/Isar.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -636,8 +636,8 @@
   thus ?thesis by simp
 qed
 
-text{*\index{cases@@{text"cases"}|(}Function @{text tl} (''tail'') is defined by @{thm tl.simps(1)} and
-@{thm tl.simps(2)}. Note that the result type of @{const length} is @{typ nat}
+text{*\index{cases@@{text"cases"}|(}Function @{text tl} (''tail'') is defined by @{thm list.sel(2)} and
+@{thm list.sel(3)}. Note that the result type of @{const length} is @{typ nat}
 and @{prop"0 - 1 = (0::nat)"}.
 
 This proof pattern works for any term @{text t} whose type is a datatype.
--- a/src/Doc/Tutorial/CTL/CTL.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/Doc/Tutorial/CTL/CTL.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -258,9 +258,9 @@
 @{thm[source]infinity_lemma}. Aficionados of minimal proofs might like to know
 that we could have given the witness without having to define a new function:
 the term
-@{term[display]"nat_rec s (\<lambda>n t. SOME u. (t,u)\<in>M \<and> Q u)"}
+@{term[display]"rec_nat s (\<lambda>n t. SOME u. (t,u)\<in>M \<and> Q u)"}
 is extensionally equal to @{term"path s Q"},
-where @{term nat_rec} is the predefined primitive recursor on @{typ nat}.
+where @{term rec_nat} is the predefined primitive recursor on @{typ nat}.
 *};
 (*<*)
 lemma
@@ -270,7 +270,7 @@
  "\<exists> p. s = p 0 \<and> (\<forall> i. (p i,p(Suc i))\<in>M \<and> Q(p i))");
  apply(simp add: Paths_def);
  apply(blast);
-apply(rule_tac x = "nat_rec s (\<lambda>n t. SOME u. (t,u)\<in>M \<and> Q u)" in exI);
+apply(rule_tac x = "rec_nat s (\<lambda>n t. SOME u. (t,u)\<in>M \<and> Q u)" in exI);
 apply(simp);
 apply(intro strip);
 apply(induct_tac i);
--- a/src/HOL/Algebra/Group.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Algebra/Group.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -34,13 +34,13 @@
 
 overloading nat_pow == "pow :: [_, 'a, nat] => 'a"
 begin
-  definition "nat_pow G a n = nat_rec \<one>\<^bsub>G\<^esub> (%u b. b \<otimes>\<^bsub>G\<^esub> a) n"
+  definition "nat_pow G a n = rec_nat \<one>\<^bsub>G\<^esub> (%u b. b \<otimes>\<^bsub>G\<^esub> a) n"
 end
 
 overloading int_pow == "pow :: [_, 'a, int] => 'a"
 begin
   definition "int_pow G a z =
-   (let p = nat_rec \<one>\<^bsub>G\<^esub> (%u b. b \<otimes>\<^bsub>G\<^esub> a)
+   (let p = rec_nat \<one>\<^bsub>G\<^esub> (%u b. b \<otimes>\<^bsub>G\<^esub> a)
     in if z < 0 then inv\<^bsub>G\<^esub> (p (nat (-z))) else p (nat z))"
 end
 
--- a/src/HOL/Auth/Guard/Extensions.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Auth/Guard/Extensions.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -410,6 +410,7 @@
 lemma knows_sub_app: "knows A evs <= knows A (evs @ evs')"
 apply (induct evs, auto)
 apply (simp add: knows_decomp)
+apply (rename_tac a b c)
 by (case_tac a, auto simp: knows.simps)
 
 subsubsection{*maximum knowledge an agent can have
@@ -504,7 +505,8 @@
 by (induct evs, auto split: event.split)
 
 lemma used'_parts [rule_format]: "X:used' evs ==> Y:parts {X} --> Y:used' evs"
-apply (induct evs, simp) 
+apply (induct evs, simp)
+apply (rename_tac a b)
 apply (case_tac a, simp_all) 
 apply (blast dest: parts_trans)+; 
 done
@@ -521,7 +523,7 @@
 by (auto dest: used_sub_Cons [THEN subsetD])
 
 lemma used_appD [dest]: "X:used (evs @ evs') ==> X:used evs | X:used evs'"
-by (induct evs, auto, case_tac a, auto)
+by (induct evs, auto, rename_tac a b, case_tac a, auto)
 
 lemma used_ConsD: "X:used (ev#evs) ==> X:used [ev] | X:used evs"
 by (case_tac ev, auto)
@@ -572,6 +574,7 @@
 apply (case_tac "A=Spy", blast)
 apply (induct evs)
 apply (simp add: used.simps, blast)
+apply (rename_tac a evs)
 apply (frule_tac ev=a and evs=evs in one_step_Cons, simp, clarify)
 apply (drule_tac P="%G. X:parts G" in knows_Cons_substD, safe)
 apply (erule initState_used)
@@ -585,6 +588,7 @@
 apply force
 apply (induct evs)
 apply (simp add: knows_max_def used.simps, blast)
+apply (rename_tac a evs)
 apply (frule_tac ev=a and evs=evs in one_step_Cons, simp, clarify)
 apply (drule_tac P="%G. X:parts G" in knows_max_Cons_substD, safe)
 apply (case_tac a, auto)
--- a/src/HOL/Auth/Guard/Guard.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Auth/Guard/Guard.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -228,6 +228,7 @@
 lemma kparts_set: "EX l'. kparts (set l) = set l' & cnb l' = cnb l"
 apply (induct l)
 apply (rule_tac x="[]" in exI, simp, clarsimp)
+apply (rename_tac a b l')
 apply (subgoal_tac "EX l''.  kparts {a} = set l'' & cnb l'' = crypt_nb a", clarify)
 apply (rule_tac x="l''@l'" in exI, simp)
 apply (rule kparts_insert_substI, simp)
--- a/src/HOL/Auth/Guard/GuardK.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Auth/Guard/GuardK.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -222,6 +222,7 @@
 lemma kparts_set: "EX l'. kparts (set l) = set l' & cnb l' = cnb l"
 apply (induct l)
 apply (rule_tac x="[]" in exI, simp, clarsimp)
+apply (rename_tac a b l')
 apply (subgoal_tac "EX l''.  kparts {a} = set l'' & cnb l'' = crypt_nb a", clarify)
 apply (rule_tac x="l''@l'" in exI, simp)
 apply (rule kparts_insert_substI, simp)
--- a/src/HOL/Auth/KerberosIV.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Auth/KerberosIV.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -256,23 +256,27 @@
 
 lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)"
 apply (induct_tac "evs")
-apply (induct_tac [2] "a", auto)
+apply (rename_tac [2] a b)
+apply (induct_tac [2] a, auto)
 done
 
 lemma spies_Gets_rev: "spies (evs @ [Gets A X]) = spies evs"
 apply (induct_tac "evs")
-apply (induct_tac [2] "a", auto)
+apply (rename_tac [2] a b)
+apply (induct_tac [2] a, auto)
 done
 
 lemma spies_Notes_rev: "spies (evs @ [Notes A X]) =
           (if A:bad then insert X (spies evs) else spies evs)"
 apply (induct_tac "evs")
-apply (induct_tac [2] "a", auto)
+apply (rename_tac [2] a b)
+apply (induct_tac [2] a, auto)
 done
 
 lemma spies_evs_rev: "spies evs = spies (rev evs)"
 apply (induct_tac "evs")
-apply (induct_tac [2] "a")
+apply (rename_tac [2] a b)
+apply (induct_tac [2] a)
 apply (simp_all (no_asm_simp) add: spies_Says_rev spies_Gets_rev spies_Notes_rev)
 done
 
@@ -280,6 +284,7 @@
 
 lemma spies_takeWhile: "spies (takeWhile P evs) <=  spies evs"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
 txt{* Resembles @{text"used_subset_append"} in theory Event.*}
 done
@@ -407,6 +412,7 @@
 lemma used_Says_rev: "used (evs @ [Says A B X]) = parts {X} \<union> (used evs)"
 apply (induct_tac "evs")
 apply simp
+apply (rename_tac a b)
 apply (induct_tac "a")
 apply auto
 done
@@ -414,6 +420,7 @@
 lemma used_Notes_rev: "used (evs @ [Notes A X]) = parts {X} \<union> (used evs)"
 apply (induct_tac "evs")
 apply simp
+apply (rename_tac a b)
 apply (induct_tac "a")
 apply auto
 done
@@ -421,6 +428,7 @@
 lemma used_Gets_rev: "used (evs @ [Gets B X]) = used evs"
 apply (induct_tac "evs")
 apply simp
+apply (rename_tac a b)
 apply (induct_tac "a")
 apply auto
 done
@@ -428,6 +436,7 @@
 lemma used_evs_rev: "used evs = used (rev evs)"
 apply (induct_tac "evs")
 apply simp
+apply (rename_tac a b)
 apply (induct_tac "a")
 apply (simp add: used_Says_rev)
 apply (simp add: used_Gets_rev)
@@ -438,6 +447,7 @@
       "x : used (takeWhile P X) --> x : used X"
 apply (induct_tac "X")
 apply simp
+apply (rename_tac a b)
 apply (induct_tac "a")
 apply (simp_all add: used_Nil)
 apply (blast dest!: initState_into_used)+
--- a/src/HOL/Auth/KerberosV.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Auth/KerberosV.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -207,22 +207,26 @@
 
 lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
 done
 
 lemma spies_Gets_rev: "spies (evs @ [Gets A X]) = spies evs"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
 done
 
 lemma spies_Notes_rev: "spies (evs @ [Notes A X]) =
           (if A:bad then insert X (spies evs) else spies evs)"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
 done
 
 lemma spies_evs_rev: "spies evs = spies (rev evs)"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a")
 apply (simp_all (no_asm_simp) add: spies_Says_rev spies_Gets_rev spies_Notes_rev)
 done
@@ -231,6 +235,7 @@
 
 lemma spies_takeWhile: "spies (takeWhile P evs) <=  spies evs"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
 txt{* Resembles @{text"used_subset_append"} in theory Event.*}
 done
--- a/src/HOL/Auth/Kerberos_BAN.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Auth/Kerberos_BAN.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -139,22 +139,26 @@
 
 lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
 done
 
 lemma spies_Gets_rev: "spies (evs @ [Gets A X]) = spies evs"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
 done
 
 lemma spies_Notes_rev: "spies (evs @ [Notes A X]) =
           (if A:bad then insert X (spies evs) else spies evs)"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
 done
 
 lemma spies_evs_rev: "spies evs = spies (rev evs)"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a")
 apply (simp_all (no_asm_simp) add: spies_Says_rev spies_Gets_rev spies_Notes_rev)
 done
@@ -163,6 +167,7 @@
 
 lemma spies_takeWhile: "spies (takeWhile P evs) <=  spies evs"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
 txt{* Resembles @{text"used_subset_append"} in theory Event.*}
 done
@@ -174,6 +179,7 @@
 lemma used_Says_rev: "used (evs @ [Says A B X]) = parts {X} \<union> (used evs)"
 apply (induct_tac "evs")
 apply simp
+apply (rename_tac a b)
 apply (induct_tac "a")
 apply auto
 done
@@ -181,6 +187,7 @@
 lemma used_Notes_rev: "used (evs @ [Notes A X]) = parts {X} \<union> (used evs)"
 apply (induct_tac "evs")
 apply simp
+apply (rename_tac a b)
 apply (induct_tac "a")
 apply auto
 done
@@ -188,6 +195,7 @@
 lemma used_Gets_rev: "used (evs @ [Gets B X]) = used evs"
 apply (induct_tac "evs")
 apply simp
+apply (rename_tac a b)
 apply (induct_tac "a")
 apply auto
 done
@@ -195,6 +203,7 @@
 lemma used_evs_rev: "used evs = used (rev evs)"
 apply (induct_tac "evs")
 apply simp
+apply (rename_tac a b)
 apply (induct_tac "a")
 apply (simp add: used_Says_rev)
 apply (simp add: used_Gets_rev)
@@ -205,6 +214,7 @@
       "x : used (takeWhile P X) --> x : used X"
 apply (induct_tac "X")
 apply simp
+apply (rename_tac a b)
 apply (induct_tac "a")
 apply (simp_all add: used_Nil)
 apply (blast dest!: initState_into_used)+
--- a/src/HOL/Auth/Kerberos_BAN_Gets.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Auth/Kerberos_BAN_Gets.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -168,6 +168,7 @@
 lemma used_Says_rev: "used (evs @ [Says A B X]) = parts {X} \<union> (used evs)"
 apply (induct_tac "evs")
 apply simp
+apply (rename_tac a b)
 apply (induct_tac "a")
 apply auto
 done
@@ -175,6 +176,7 @@
 lemma used_Notes_rev: "used (evs @ [Notes A X]) = parts {X} \<union> (used evs)"
 apply (induct_tac "evs")
 apply simp
+apply (rename_tac a b)
 apply (induct_tac "a")
 apply auto
 done
@@ -182,6 +184,7 @@
 lemma used_Gets_rev: "used (evs @ [Gets B X]) = used evs"
 apply (induct_tac "evs")
 apply simp
+apply (rename_tac a b)
 apply (induct_tac "a")
 apply auto
 done
@@ -189,6 +192,7 @@
 lemma used_evs_rev: "used evs = used (rev evs)"
 apply (induct_tac "evs")
 apply simp
+apply (rename_tac a b)
 apply (induct_tac "a")
 apply (simp add: used_Says_rev)
 apply (simp add: used_Gets_rev)
@@ -199,6 +203,7 @@
       "x : used (takeWhile P X) --> x : used X"
 apply (induct_tac "X")
 apply simp
+apply (rename_tac a b)
 apply (induct_tac "a")
 apply (simp_all add: used_Nil)
 apply (blast dest!: initState_into_used)+
--- a/src/HOL/Auth/NS_Shared.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Auth/NS_Shared.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -385,22 +385,26 @@
 
 lemma spies_Says_rev: "spies (evs @ [Says A B X]) = insert X (spies evs)"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
 done
 
 lemma spies_Gets_rev: "spies (evs @ [Gets A X]) = spies evs"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
 done
 
 lemma spies_Notes_rev: "spies (evs @ [Notes A X]) =
           (if A:bad then insert X (spies evs) else spies evs)"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
 done
 
 lemma spies_evs_rev: "spies evs = spies (rev evs)"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a")
 apply (simp_all (no_asm_simp) add: spies_Says_rev spies_Gets_rev spies_Notes_rev)
 done
@@ -409,6 +413,7 @@
 
 lemma spies_takeWhile: "spies (takeWhile P evs) <=  spies evs"
 apply (induct_tac "evs")
+apply (rename_tac [2] a b)
 apply (induct_tac [2] "a", auto)
 txt{* Resembles @{text"used_subset_append"} in theory Event.*}
 done
--- a/src/HOL/Auth/Public.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Auth/Public.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -61,7 +61,7 @@
   injective_publicKey:
     "publicKey b A = publicKey c A' ==> b=c & A=A'"
    apply (rule exI [of _ 
-       "%b A. 2 * agent_case 0 (\<lambda>n. n + 2) 1 A + keymode_case 0 1 b"])
+       "%b A. 2 * case_agent 0 (\<lambda>n. n + 2) 1 A + case_keymode 0 1 b"])
    apply (auto simp add: inj_on_def split: agent.split keymode.split)
    apply presburger
    apply presburger
@@ -137,7 +137,7 @@
 specification (shrK)
   inj_shrK: "inj shrK"
   --{*No two agents have the same long-term key*}
-   apply (rule exI [of _ "agent_case 0 (\<lambda>n. n + 2) 1"]) 
+   apply (rule exI [of _ "case_agent 0 (\<lambda>n. n + 2) 1"]) 
    apply (simp add: inj_on_def split: agent.split) 
    done
 
--- a/src/HOL/Auth/Shared.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Auth/Shared.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -17,7 +17,7 @@
 specification (shrK)
   inj_shrK: "inj shrK"
   --{*No two agents have the same long-term key*}
-   apply (rule exI [of _ "agent_case 0 (\<lambda>n. n + 2) 1"]) 
+   apply (rule exI [of _ "case_agent 0 (\<lambda>n. n + 2) 1"]) 
    apply (simp add: inj_on_def split: agent.split) 
    done
 
--- a/src/HOL/Auth/TLS.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Auth/TLS.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -81,7 +81,7 @@
   inj_sessionK: "inj sessionK"
   --{*sessionK is collision-free; also, no clientK clashes with any serverK.*}
    apply (rule exI [of _ 
-         "%((x,y,z), r). prod_encode(role_case 0 1 r, 
+         "%((x,y,z), r). prod_encode(case_role 0 1 r, 
                            prod_encode(x, prod_encode(y,z)))"])
    apply (simp add: inj_on_def prod_encode_eq split: role.split) 
    done
--- a/src/HOL/BNF_Def.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/BNF_Def.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -118,9 +118,9 @@
 lemma predicate2_eqD: "A = B \<Longrightarrow> A a b \<longleftrightarrow> B a b"
 by metis
 
-lemma sum_case_o_inj:
-"sum_case f g \<circ> Inl = f"
-"sum_case f g \<circ> Inr = g"
+lemma case_sum_o_inj:
+"case_sum f g \<circ> Inl = f"
+"case_sum f g \<circ> Inr = g"
 by auto
 
 lemma card_order_csum_cone_cexp_def:
--- a/src/HOL/BNF_Examples/Derivation_Trees/Gram_Lang.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/BNF_Examples/Derivation_Trees/Gram_Lang.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -860,7 +860,7 @@
   have "subtr (insert n ?ns1) (f (last (n1 # nl))) (f n)"
   using f subtr.Step[OF _ fn1_flast fn1] by auto
   thus ?case unfolding 1 by simp
-qed (metis f hd.simps last_ConsL last_in_set not_Cons_self2 subtr.Refl)
+qed (metis f list.sel(1) last_ConsL last_in_set not_Cons_self2 subtr.Refl)
 
 lemma reg_subtr_path_aux:
 assumes f: "reg f tr" and n: "subtr ns tr1 tr"
@@ -878,7 +878,7 @@
   obtain nl where nl: "path f nl" and f_nl: "f (hd nl) = tr1"
   and last_nl: "f (last nl) = tr2" and set: "set nl \<subseteq> ns" using Step(3)[OF tr1] by auto
   have 0: "path f (root tr # nl)" apply (subst path.simps)
-  using f_nl nl reg_root tr tr1_tr by (metis hd.simps neq_Nil_conv)
+  using f_nl nl reg_root tr tr1_tr by (metis list.sel(1) neq_Nil_conv)
   show ?case apply(rule exI[of _ "(root tr) # nl"])
   using 0 reg_root tr last_nl nl path_NE rtr set by auto
 qed
--- a/src/HOL/BNF_Examples/Derivation_Trees/Prelim.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/BNF_Examples/Derivation_Trees/Prelim.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -27,8 +27,8 @@
 lemma sum_map_InrD: "(f \<oplus> g) z = Inr x \<Longrightarrow> \<exists>y. z = Inr y \<and> g y = x"
 by (cases z) auto
 
-abbreviation sum_case_abbrev ("[[_,_]]" 800)
-where "[[f,g]] \<equiv> Sum_Type.sum_case f g"
+abbreviation case_sum_abbrev ("[[_,_]]" 800)
+where "[[f,g]] \<equiv> Sum_Type.case_sum f g"
 
 lemma Inl_oplus_elim:
 assumes "Inl tr \<in> (id \<oplus> f) ` tns"
--- a/src/HOL/BNF_FP_Base.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/BNF_FP_Base.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -19,10 +19,10 @@
 lemma eq_sym_Unity_conv: "(x = (() = ())) = x"
 by blast
 
-lemma unit_case_Unity: "(case u of () \<Rightarrow> f) = f"
+lemma case_unit_Unity: "(case u of () \<Rightarrow> f) = f"
 by (cases u) (hypsubst, rule unit.cases)
 
-lemma prod_case_Pair_iden: "(case p of (x, y) \<Rightarrow> (x, y)) = p"
+lemma case_prod_Pair_iden: "(case p of (x, y) \<Rightarrow> (x, y)) = p"
 by simp
 
 lemma unit_all_impI: "(P () \<Longrightarrow> Q ()) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
@@ -53,9 +53,9 @@
 
 lemma ssubst_mem: "\<lbrakk>t = s; s \<in> X\<rbrakk> \<Longrightarrow> t \<in> X" by simp
 
-lemma sum_case_step:
-"sum_case (sum_case f' g') g (Inl p) = sum_case f' g' p"
-"sum_case f (sum_case f' g') (Inr p) = sum_case f' g' p"
+lemma case_sum_step:
+"case_sum (case_sum f' g') g (Inl p) = case_sum f' g' p"
+"case_sum f (case_sum f' g') (Inr p) = case_sum f' g' p"
 by auto
 
 lemma one_pointE: "\<lbrakk>\<And>x. s = x \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
@@ -71,8 +71,8 @@
 lemma obj_sumE: "\<lbrakk>\<forall>x. s = Inl x \<longrightarrow> P; \<forall>x. s = Inr x \<longrightarrow> P\<rbrakk> \<Longrightarrow> P"
 by (cases s) auto
 
-lemma sum_case_if:
-"sum_case f g (if p then Inl x else Inr y) = (if p then f x else g y)"
+lemma case_sum_if:
+"case_sum f g (if p then Inl x else Inr y) = (if p then f x else g y)"
 by simp
 
 lemma mem_UN_compreh_eq: "(z : \<Union>{y. \<exists>x\<in>A. y = F x}) = (\<exists>x\<in>A. z : F x)"
@@ -122,14 +122,14 @@
 lemma map_pair_o_convol_id: "(map_pair f id \<circ> <id , g>) x = <id \<circ> f , g> x"
   unfolding map_pair_o_convol id_comp comp_id ..
 
-lemma o_sum_case: "h o sum_case f g = sum_case (h o f) (h o g)"
+lemma o_case_sum: "h o case_sum f g = case_sum (h o f) (h o g)"
   unfolding comp_def by (auto split: sum.splits)
 
-lemma sum_case_o_sum_map: "sum_case f g o sum_map h1 h2 = sum_case (f o h1) (g o h2)"
+lemma case_sum_o_sum_map: "case_sum f g o sum_map h1 h2 = case_sum (f o h1) (g o h2)"
   unfolding comp_def by (auto split: sum.splits)
 
-lemma sum_case_o_sum_map_id: "(sum_case id g o sum_map f id) x = sum_case (f o id) g x"
-  unfolding sum_case_o_sum_map id_comp comp_id ..
+lemma case_sum_o_sum_map_id: "(case_sum id g o sum_map f id) x = case_sum (f o id) g x"
+  unfolding case_sum_o_sum_map id_comp comp_id ..
 
 lemma fun_rel_def_butlast:
   "(fun_rel R (fun_rel S T)) f g = (\<forall>x y. R x y \<longrightarrow> (fun_rel S T) (f x) (g y))"
--- a/src/HOL/BNF_GFP.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/BNF_GFP.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -27,13 +27,13 @@
 lemma neq_eq_eq_contradict: "\<lbrakk>t \<noteq> u; s = t; s = u\<rbrakk> \<Longrightarrow> P"
 by fast
 
-lemma sum_case_expand_Inr: "f o Inl = g \<Longrightarrow> f x = sum_case g (f o Inr) x"
+lemma case_sum_expand_Inr: "f o Inl = g \<Longrightarrow> f x = case_sum g (f o Inr) x"
 by (auto split: sum.splits)
 
-lemma sum_case_expand_Inr': "f o Inl = g \<Longrightarrow> h = f o Inr \<longleftrightarrow> sum_case g h = f"
+lemma case_sum_expand_Inr': "f o Inl = g \<Longrightarrow> h = f o Inr \<longleftrightarrow> case_sum g h = f"
 apply rule
  apply (rule ext, force split: sum.split)
-by (rule ext, metis sum_case_o_inj(2))
+by (rule ext, metis case_sum_o_inj(2))
 
 lemma converse_Times: "(A \<times> B) ^-1 = B \<times> A"
 by fast
@@ -266,16 +266,16 @@
 lemma Inr_Field_csum: "a \<in> Field s \<Longrightarrow> Inr a \<in> Field (r +c s)"
 unfolding Field_card_of csum_def by auto
 
-lemma nat_rec_0_imp: "f = nat_rec f1 (%n rec. f2 n rec) \<Longrightarrow> f 0 = f1"
+lemma rec_nat_0_imp: "f = rec_nat f1 (%n rec. f2 n rec) \<Longrightarrow> f 0 = f1"
 by auto
 
-lemma nat_rec_Suc_imp: "f = nat_rec f1 (%n rec. f2 n rec) \<Longrightarrow> f (Suc n) = f2 n (f n)"
+lemma rec_nat_Suc_imp: "f = rec_nat f1 (%n rec. f2 n rec) \<Longrightarrow> f (Suc n) = f2 n (f n)"
 by auto
 
-lemma list_rec_Nil_imp: "f = list_rec f1 (%x xs rec. f2 x xs rec) \<Longrightarrow> f [] = f1"
+lemma rec_list_Nil_imp: "f = rec_list f1 (%x xs rec. f2 x xs rec) \<Longrightarrow> f [] = f1"
 by auto
 
-lemma list_rec_Cons_imp: "f = list_rec f1 (%x xs rec. f2 x xs rec) \<Longrightarrow> f (x # xs) = f2 x xs (f xs)"
+lemma rec_list_Cons_imp: "f = rec_list f1 (%x xs rec. f2 x xs rec) \<Longrightarrow> f (x # xs) = f2 x xs (f xs)"
 by auto
 
 lemma not_arg_cong_Inr: "x \<noteq> y \<Longrightarrow> Inr x \<noteq> Inr y"
--- a/src/HOL/Bali/Basis.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Bali/Basis.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -147,7 +147,7 @@
 
 hide_const In0 In1
 
-notation sum_case  (infixr "'(+')"80)
+notation case_sum  (infixr "'(+')"80)
 
 primrec the_Inl :: "'a + 'b \<Rightarrow> 'a"
   where "the_Inl (Inl a) = a"
--- a/src/HOL/Bali/Conform.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Bali/Conform.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -437,9 +437,10 @@
 lemma conforms_absorb [rule_format]:
   "(a, b)\<Colon>\<preceq>(G, L) \<longrightarrow> (absorb j a, b)\<Colon>\<preceq>(G, L)"
 apply (rule impI)
-apply ( case_tac a)
+apply (case_tac a)
 apply (case_tac "absorb j a")
 apply auto
+apply (rename_tac a)
 apply (case_tac "absorb j (Some a)",auto)
 apply (erule conforms_NormI)
 done
@@ -554,5 +555,4 @@
 apply (force dest: conforms_globsD)+
 done
 
-
 end
--- a/src/HOL/Bali/Eval.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Bali/Eval.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -141,7 +141,7 @@
   where "\<lfloor>es\<rfloor>\<^sub>l == In3 es"
 
 definition undefined3 :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> vals" where
- "undefined3 = sum3_case (In1 \<circ> sum_case (\<lambda>x. undefined) (\<lambda>x. Unit))
+ "undefined3 = case_sum3 (In1 \<circ> case_sum (\<lambda>x. undefined) (\<lambda>x. Unit))
                      (\<lambda>x. In2 undefined) (\<lambda>x. In3 undefined)"
 
 lemma [simp]: "undefined3 (In1l x) = In1 undefined"
--- a/src/HOL/Bali/State.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Bali/State.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -232,11 +232,11 @@
 
 definition
   globs :: "st \<Rightarrow> globs"
-  where "globs = st_case (\<lambda>g l. g)"
+  where "globs = case_st (\<lambda>g l. g)"
   
 definition
   locals :: "st \<Rightarrow> locals"
-  where "locals = st_case (\<lambda>g l. l)"
+  where "locals = case_st (\<lambda>g l. l)"
 
 definition heap :: "st \<Rightarrow> heap" where
  "heap s = globs s \<circ> Heap"
@@ -303,19 +303,19 @@
 
 definition
   gupd :: "oref  \<Rightarrow> obj \<Rightarrow> st \<Rightarrow> st" ("gupd'(_\<mapsto>_')" [10, 10] 1000)
-  where "gupd r obj = st_case (\<lambda>g l. st (g(r\<mapsto>obj)) l)"
+  where "gupd r obj = case_st (\<lambda>g l. st (g(r\<mapsto>obj)) l)"
 
 definition
   lupd :: "lname \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st" ("lupd'(_\<mapsto>_')" [10, 10] 1000)
-  where "lupd vn v = st_case (\<lambda>g l. st g (l(vn\<mapsto>v)))"
+  where "lupd vn v = case_st (\<lambda>g l. st g (l(vn\<mapsto>v)))"
 
 definition
   upd_gobj :: "oref \<Rightarrow> vn \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st"
-  where "upd_gobj r n v = st_case (\<lambda>g l. st (chg_map (upd_obj n v) r g) l)"
+  where "upd_gobj r n v = case_st (\<lambda>g l. st (chg_map (upd_obj n v) r g) l)"
 
 definition
   set_locals  :: "locals \<Rightarrow> st \<Rightarrow> st"
-  where "set_locals l = st_case (\<lambda>g l'. st g l)"
+  where "set_locals l = case_st (\<lambda>g l'. st g l)"
 
 definition
   init_obj :: "prog \<Rightarrow> obj_tag \<Rightarrow> oref \<Rightarrow> st \<Rightarrow> st"
--- a/src/HOL/Bali/TypeSafe.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Bali/TypeSafe.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -825,7 +825,7 @@
 
 
 lemma lconf_map_lname [simp]: 
-  "G,s\<turnstile>(lname_case l1 l2)[\<Colon>\<preceq>](lname_case L1 L2)
+  "G,s\<turnstile>(case_lname l1 l2)[\<Colon>\<preceq>](case_lname L1 L2)
    =
   (G,s\<turnstile>l1[\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>(\<lambda>x::unit . l2)[\<Colon>\<preceq>](\<lambda>x::unit. L2))"
 apply (unfold lconf_def)
@@ -833,7 +833,7 @@
 done
 
 lemma wlconf_map_lname [simp]: 
-  "G,s\<turnstile>(lname_case l1 l2)[\<sim>\<Colon>\<preceq>](lname_case L1 L2)
+  "G,s\<turnstile>(case_lname l1 l2)[\<sim>\<Colon>\<preceq>](case_lname L1 L2)
    =
   (G,s\<turnstile>l1[\<sim>\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>(\<lambda>x::unit . l2)[\<sim>\<Colon>\<preceq>](\<lambda>x::unit. L2))"
 apply (unfold wlconf_def)
@@ -841,7 +841,7 @@
 done
 
 lemma lconf_map_ename [simp]:
-  "G,s\<turnstile>(ename_case l1 l2)[\<Colon>\<preceq>](ename_case L1 L2)
+  "G,s\<turnstile>(case_ename l1 l2)[\<Colon>\<preceq>](case_ename L1 L2)
    =
   (G,s\<turnstile>l1[\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>(\<lambda>x::unit. l2)[\<Colon>\<preceq>](\<lambda>x::unit. L2))"
 apply (unfold lconf_def)
@@ -849,7 +849,7 @@
 done
 
 lemma wlconf_map_ename [simp]:
-  "G,s\<turnstile>(ename_case l1 l2)[\<sim>\<Colon>\<preceq>](ename_case L1 L2)
+  "G,s\<turnstile>(case_ename l1 l2)[\<sim>\<Colon>\<preceq>](case_ename L1 L2)
    =
   (G,s\<turnstile>l1[\<sim>\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>(\<lambda>x::unit. l2)[\<sim>\<Colon>\<preceq>](\<lambda>x::unit. L2))"
 apply (unfold wlconf_def)
@@ -1436,9 +1436,9 @@
 
    
 lemma dom_vname_split:
- "dom (lname_case (ename_case (tab(x\<mapsto>y)(xs[\<mapsto>]ys)) a) b)
-   = dom (lname_case (ename_case (tab(x\<mapsto>y)) a) b) \<union> 
-     dom (lname_case (ename_case (tab(xs[\<mapsto>]ys)) a) b)"
+ "dom (case_lname (case_ename (tab(x\<mapsto>y)(xs[\<mapsto>]ys)) a) b)
+   = dom (case_lname (case_ename (tab(x\<mapsto>y)) a) b) \<union> 
+     dom (case_lname (case_ename (tab(xs[\<mapsto>]ys)) a) b)"
   (is "?List x xs y ys = ?Hd x y \<union> ?Tl xs ys")
 proof 
   show "?List x xs y ys \<subseteq> ?Hd x y \<union> ?Tl xs ys"
@@ -1514,37 +1514,37 @@
   qed
 qed
  
-lemma dom_ename_case_None_simp:
- "dom (ename_case vname_tab None) = VNam ` (dom vname_tab)"
+lemma dom_case_ename_None_simp:
+ "dom (case_ename vname_tab None) = VNam ` (dom vname_tab)"
   apply (auto simp add: dom_def image_def )
   apply (case_tac "x")
   apply auto
   done
 
-lemma dom_ename_case_Some_simp:
- "dom (ename_case vname_tab (Some a)) = VNam ` (dom vname_tab) \<union> {Res}"
+lemma dom_case_ename_Some_simp:
+ "dom (case_ename vname_tab (Some a)) = VNam ` (dom vname_tab) \<union> {Res}"
   apply (auto simp add: dom_def image_def )
   apply (case_tac "x")
   apply auto
   done
 
-lemma dom_lname_case_None_simp:
-  "dom (lname_case ename_tab None) = EName ` (dom ename_tab)"
+lemma dom_case_lname_None_simp:
+  "dom (case_lname ename_tab None) = EName ` (dom ename_tab)"
   apply (auto simp add: dom_def image_def )
   apply (case_tac "x")
   apply auto
   done
 
-lemma dom_lname_case_Some_simp:
- "dom (lname_case ename_tab (Some a)) = EName ` (dom ename_tab) \<union> {This}"
+lemma dom_case_lname_Some_simp:
+ "dom (case_lname ename_tab (Some a)) = EName ` (dom ename_tab) \<union> {This}"
   apply (auto simp add: dom_def image_def)
   apply (case_tac "x")
   apply auto
   done
 
-lemmas dom_lname_ename_case_simps =  
-     dom_ename_case_None_simp dom_ename_case_Some_simp 
-     dom_lname_case_None_simp dom_lname_case_Some_simp
+lemmas dom_lname_case_ename_simps =  
+     dom_case_ename_None_simp dom_case_ename_Some_simp 
+     dom_case_lname_None_simp dom_case_lname_Some_simp
 
 lemma image_comp: 
  "f ` g ` A = (f \<circ> g) ` A"
@@ -1569,13 +1569,13 @@
     with static_m' dom_vnames m
     show ?thesis
       by (cases s) (simp add: init_lvars_def Let_def parameters_def
-                              dom_lname_ename_case_simps image_comp)
+                              dom_lname_case_ename_simps image_comp)
   next
     case False
     with static_m' dom_vnames m
     show ?thesis
       by (cases s) (simp add: init_lvars_def Let_def parameters_def
-                              dom_lname_ename_case_simps image_comp)
+                              dom_lname_case_ename_simps image_comp)
   qed
 qed
 
--- a/src/HOL/Cardinals/Ordinal_Arithmetic.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Cardinals/Ordinal_Arithmetic.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -365,9 +365,9 @@
 proof safe
   assume ?L
   from `?L` show ?R1 unfolding fin_support_def support_def
-    by (fastforce simp: image_iff elim: finite_surj[of _ _ "sum_case id undefined"])
+    by (fastforce simp: image_iff elim: finite_surj[of _ _ "case_sum id undefined"])
   from `?L` show ?R2 unfolding fin_support_def support_def
-    by (fastforce simp: image_iff elim: finite_surj[of _ _ "sum_case undefined id"])
+    by (fastforce simp: image_iff elim: finite_surj[of _ _ "case_sum undefined id"])
 next
   assume ?R1 ?R2
   thus ?L unfolding fin_support_def support_Un
@@ -1506,15 +1506,15 @@
 qed
 
 lemma max_fun_diff_eq_Inl:
-  assumes "wo_rel.max_fun_diff (s +o t) (sum_case f1 g1) (sum_case f2 g2) = Inl x"
-    "sum_case f1 g1 \<noteq> sum_case f2 g2"
-    "sum_case f1 g1 \<in> FinFunc r (s +o t)" "sum_case f2 g2 \<in> FinFunc r (s +o t)"
+  assumes "wo_rel.max_fun_diff (s +o t) (case_sum f1 g1) (case_sum f2 g2) = Inl x"
+    "case_sum f1 g1 \<noteq> case_sum f2 g2"
+    "case_sum f1 g1 \<in> FinFunc r (s +o t)" "case_sum f2 g2 \<in> FinFunc r (s +o t)"
   shows "wo_rel.max_fun_diff s f1 f2 = x" (is ?P) "g1 = g2" (is ?Q)
 proof -
   interpret st!: wo_rel "s +o t" by unfold_locales (rule osum_Well_order[OF s t])
   interpret s!: wo_rel s by unfold_locales (rule s)
   interpret rst!: wo_rel2 r "s +o t" by unfold_locales (rule r, rule osum_Well_order[OF s t])
-  from assms(1) have *: "st.isMaxim {a \<in> Field (s +o t). sum_case f1 g1 a \<noteq> sum_case f2 g2 a} (Inl x)"
+  from assms(1) have *: "st.isMaxim {a \<in> Field (s +o t). case_sum f1 g1 a \<noteq> case_sum f2 g2 a} (Inl x)"
     using rst.isMaxim_max_fun_diff[OF assms(2-4)] by simp
   hence "s.isMaxim {a \<in> Field s. f1 a \<noteq> f2 a} x"
     unfolding st.isMaxim_def s.isMaxim_def Field_osum by (auto simp: osum_def)
@@ -1530,15 +1530,15 @@
 qed
 
 lemma max_fun_diff_eq_Inr:
-  assumes "wo_rel.max_fun_diff (s +o t) (sum_case f1 g1) (sum_case f2 g2) = Inr x"
-    "sum_case f1 g1 \<noteq> sum_case f2 g2"
-    "sum_case f1 g1 \<in> FinFunc r (s +o t)" "sum_case f2 g2 \<in> FinFunc r (s +o t)"
+  assumes "wo_rel.max_fun_diff (s +o t) (case_sum f1 g1) (case_sum f2 g2) = Inr x"
+    "case_sum f1 g1 \<noteq> case_sum f2 g2"
+    "case_sum f1 g1 \<in> FinFunc r (s +o t)" "case_sum f2 g2 \<in> FinFunc r (s +o t)"
   shows "wo_rel.max_fun_diff t g1 g2 = x" (is ?P) "g1 \<noteq> g2" (is ?Q)
 proof -
   interpret st!: wo_rel "s +o t" by unfold_locales (rule osum_Well_order[OF s t])
   interpret t!: wo_rel t by unfold_locales (rule t)
   interpret rst!: wo_rel2 r "s +o t" by unfold_locales (rule r, rule osum_Well_order[OF s t])
-  from assms(1) have *: "st.isMaxim {a \<in> Field (s +o t). sum_case f1 g1 a \<noteq> sum_case f2 g2 a} (Inr x)"
+  from assms(1) have *: "st.isMaxim {a \<in> Field (s +o t). case_sum f1 g1 a \<noteq> case_sum f2 g2 a} (Inr x)"
     using rst.isMaxim_max_fun_diff[OF assms(2-4)] by simp
   hence "t.isMaxim {a \<in> Field t. g1 a \<noteq> g2 a} x"
     unfolding st.isMaxim_def t.isMaxim_def Field_osum by (auto simp: osum_def)
@@ -1551,7 +1551,7 @@
   interpret rst!: wo_rel2 r "s +o t" by unfold_locales (rule r, rule osum_Well_order[OF s t])
   interpret rs!: wo_rel2 r s by unfold_locales (rule r, rule s)
   interpret rt!: wo_rel2 r t by unfold_locales (rule r, rule t)
-  let ?f = "\<lambda>(f, g). sum_case f g"
+  let ?f = "\<lambda>(f, g). case_sum f g"
   have "bij_betw ?f (Field ?L) (Field ?R)"
   unfolding bij_betw_def rst.Field_oexp rs.Field_oexp rt.Field_oexp Field_oprod proof (intro conjI)
     show "inj_on ?f (FinFunc r s \<times> FinFunc r t)" unfolding inj_on_def
--- a/src/HOL/Code_Numeral.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Code_Numeral.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -384,7 +384,7 @@
     by (auto simp add: sgn_if)
   have aux2: "\<And>q::int. - int_of_integer k = int_of_integer l * q \<longleftrightarrow> int_of_integer k = int_of_integer l * - q" by auto
   show ?thesis
-    by (simp add: prod_eq_iff integer_eq_iff prod_case_beta aux1)
+    by (simp add: prod_eq_iff integer_eq_iff case_prod_beta aux1)
       (auto simp add: zdiv_zminus1_eq_if zmod_zminus1_eq_if div_minus_right mod_minus_right aux2)
 qed
 
@@ -475,7 +475,7 @@
   }
   note aux = this
   show ?thesis
-    by (auto simp add: num_of_integer_def nat_of_integer_def Let_def prod_case_beta
+    by (auto simp add: num_of_integer_def nat_of_integer_def Let_def case_prod_beta
       not_le integer_eq_iff less_eq_integer_def
       nat_mult_distrib nat_div_distrib num_of_nat_One num_of_nat_plus_distrib
        mult_2 [where 'a=nat] aux add_One)
@@ -792,7 +792,7 @@
   by (rule is_measure_trivial)
 
 
-subsection {* Inductive represenation of target language naturals *}
+subsection {* Inductive representation of target language naturals *}
 
 lift_definition Suc :: "natural \<Rightarrow> natural"
   is Nat.Suc
@@ -803,7 +803,7 @@
 rep_datatype "0::natural" Suc
   by (transfer, fact nat.induct nat.inject nat.distinct)+
 
-lemma natural_case [case_names nat, cases type: natural]:
+lemma natural_cases [case_names nat, cases type: natural]:
   fixes m :: natural
   assumes "\<And>n. m = of_nat n \<Longrightarrow> P"
   shows P
@@ -885,7 +885,7 @@
   by transfer (simp add: fun_eq_iff)
 
 lemma [code, code_unfold]:
-  "natural_case f g n = (if n = 0 then f else g (n - 1))"
+  "case_natural f g n = (if n = 0 then f else g (n - 1))"
   by (cases n rule: natural.exhaust) (simp_all, simp add: Suc_def)
 
 declare natural.recs [code del]
--- a/src/HOL/Datatype.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Datatype.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -56,7 +56,7 @@
   Push_Node_def:  "Push_Node == (%n x. Abs_Node (apfst (Push n) (Rep_Node x)))"
 
   (*crude "lists" of nats -- needed for the constructions*)
-  Push_def:   "Push == (%b h. nat_case b h)"
+  Push_def:   "Push == (%b h. case_nat b h)"
 
   (** operations on S-expressions -- sets of nodes **)
 
@@ -133,7 +133,7 @@
 
 lemma Node_Push_I: "p: Node ==> apfst (Push i) p : Node"
 apply (simp add: Node_def Push_def) 
-apply (fast intro!: apfst_conv nat_case_Suc [THEN trans])
+apply (fast intro!: apfst_conv nat.cases(2)[THEN trans])
 done
 
 
@@ -251,7 +251,7 @@
 by (simp add: ndepth_def  Node_K0_I [THEN Abs_Node_inverse] Least_equality)
 
 lemma ndepth_Push_Node_aux:
-     "nat_case (Inr (Suc i)) f k = Inr 0 --> Suc(LEAST x. f x = Inr 0) <= k"
+     "case_nat (Inr (Suc i)) f k = Inr 0 --> Suc(LEAST x. f x = Inr 0) <= k"
 apply (induct_tac "k", auto)
 apply (erule Least_le)
 done
--- a/src/HOL/Decision_Procs/Approximation.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Decision_Procs/Approximation.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -182,7 +182,7 @@
               else if x < 0 then - ub_sqrt prec (- x)
                             else 0)"
 by pat_completeness auto
-termination by (relation "measure (\<lambda> v. let (prec, x) = sum_case id id v in (if x < 0 then 1 else 0))", auto)
+termination by (relation "measure (\<lambda> v. let (prec, x) = case_sum id id v in (if x < 0 then 1 else 0))", auto)
 
 declare lb_sqrt.simps[simp del]
 declare ub_sqrt.simps[simp del]
@@ -484,7 +484,7 @@
                                            else Float 1 1 * ub_horner y
                           else ub_pi prec * Float 1 -1 - lb_horner (float_divl prec 1 x)))"
 by pat_completeness auto
-termination by (relation "measure (\<lambda> v. let (prec, x) = sum_case id id v in (if x < 0 then 1 else 0))", auto)
+termination by (relation "measure (\<lambda> v. let (prec, x) = case_sum id id v in (if x < 0 then 1 else 0))", auto)
 
 declare ub_arctan_horner.simps[simp del]
 declare lb_arctan_horner.simps[simp del]
@@ -1387,7 +1387,7 @@
              else if x < - 1  then ub_exp_horner prec (get_odd (prec + 2)) 1 1 (float_divr prec x (- floor_fl x)) ^ (nat (- int_floor_fl x))
                               else ub_exp_horner prec (get_odd (prec + 2)) 1 1 x)"
 by pat_completeness auto
-termination by (relation "measure (\<lambda> v. let (prec, x) = sum_case id id v in (if 0 < x then 1 else 0))", auto)
+termination by (relation "measure (\<lambda> v. let (prec, x) = case_sum id id v in (if 0 < x then 1 else 0))", auto)
 
 lemma exp_m1_ge_quarter: "(1 / 4 :: real) \<le> exp (- 1)"
 proof -
@@ -1709,7 +1709,7 @@
                                         Some (lb_ln2 prec * (Float (exponent x + l) 0) + horner (Float (mantissa x) (- l) - 1)))"
 by pat_completeness auto
 
-termination proof (relation "measure (\<lambda> v. let (prec, x) = sum_case id id v in (if x < 1 then 1 else 0))", auto)
+termination proof (relation "measure (\<lambda> v. let (prec, x) = case_sum id id v in (if x < 1 then 1 else 0))", auto)
   fix prec and x :: float assume "\<not> real x \<le> 0" and "real x < 1" and "real (float_divl (max prec (Suc 0)) 1 x) < 1"
   hence "0 < real x" "1 \<le> max prec (Suc 0)" "real x < 1" by auto
   from float_divl_pos_less1_bound[OF `0 < real x` `real x < 1` `1 \<le> max prec (Suc 0)`]
@@ -3062,7 +3062,7 @@
   case 0
   then obtain ly uy
     where *: "approx_tse prec 0 t ((l + u) * Float 1 -1) 1 f [Some (l, u)] = Some (ly, uy)"
-    and **: "cmp ly uy" by (auto elim!: option_caseE)
+    and **: "cmp ly uy" by (auto elim!: case_optionE)
   with 0 show ?case by auto
 next
   case (Suc s)
@@ -3163,7 +3163,7 @@
   with assms obtain l u l' u'
     where a: "approx prec a [None] = Some (l, u)"
     and b: "approx prec b [None] = Some (l', u')"
-    unfolding approx_tse_form_def by (auto elim!: option_caseE)
+    unfolding approx_tse_form_def by (auto elim!: case_optionE)
 
   from Bound assms have "i = Var 0" unfolding approx_tse_form_def by auto
   hence i: "interpret_floatarith i [x] = x" by auto
@@ -3198,10 +3198,10 @@
       show ?thesis using AtLeastAtMost by auto
     next
       case (Bound x a b f') with assms
-      show ?thesis by (auto elim!: option_caseE simp add: f_def approx_tse_form_def)
+      show ?thesis by (auto elim!: case_optionE simp add: f_def approx_tse_form_def)
     next
       case (Assign x a f') with assms
-      show ?thesis by (auto elim!: option_caseE simp add: f_def approx_tse_form_def)
+      show ?thesis by (auto elim!: case_optionE simp add: f_def approx_tse_form_def)
     qed } thus ?thesis unfolding f_def by auto
 next
   case Assign
--- a/src/HOL/Decision_Procs/Ferrack.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Decision_Procs/Ferrack.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -2030,4 +2030,3 @@
   by rferrack
 
 end
-
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -3003,5 +3003,3 @@
 oops
 *)
 end
-
-
--- a/src/HOL/Decision_Procs/Polynomial_List.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Decision_Procs/Polynomial_List.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -919,6 +919,7 @@
   "last ((a %* p) +++ (x#(b %* p))) = (if p = [] then x else b * last p)"
   apply (induct p arbitrary: a x b)
   apply auto
+  apply (rename_tac a p aa x b)
   apply (subgoal_tac "padd (cmult aa p) (times b a # cmult b p) \<noteq> []")
   apply simp
   apply (induct_tac p)
@@ -1042,6 +1043,7 @@
 lemma poly_mono: "abs(x) \<le> k \<Longrightarrow> abs(poly p (x::'a::{linordered_idom})) \<le> poly (map abs p) k"
   apply (induct p)
   apply auto
+  apply (rename_tac a p)
   apply (rule_tac y = "abs a + abs (x * poly p x)" in order_trans)
   apply (rule abs_triangle_ineq)
   apply (auto intro!: mult_mono simp add: abs_mult)
--- a/src/HOL/Divides.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Divides.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -928,7 +928,7 @@
 
 lemma divmod_nat_if [code]: "divmod_nat m n = (if n = 0 \<or> m < n then (0, m) else
   let (q, r) = divmod_nat (m - n) n in (Suc q, r))"
-  by (simp add: prod_eq_iff prod_case_beta not_less le_div_geq le_mod_geq)
+  by (simp add: prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq)
 
 text {* Simproc for cancelling @{const div} and @{const mod} *}
 
--- a/src/HOL/Extraction.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Extraction.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -5,7 +5,7 @@
 header {* Program extraction for HOL *}
 
 theory Extraction
-imports Option
+imports Datatype Option
 begin
 
 ML_file "Tools/rewrite_hol_proof.ML"
--- a/src/HOL/Fun.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Fun.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -662,10 +662,10 @@
   "_Update f (_updbinds b bs)" == "_Update (_Update f b) bs"
   "f(x:=y)" == "CONST fun_upd f x y"
 
-(* Hint: to define the sum of two functions (or maps), use sum_case.
+(* Hint: to define the sum of two functions (or maps), use case_sum.
          A nice infix syntax could be defined (in Datatype.thy or below) by
 notation
-  sum_case  (infixr "'(+')"80)
+  case_sum  (infixr "'(+')"80)
 *)
 
 lemma fun_upd_idem_iff: "(f(x:=y) = f) = (f x = y)"
--- a/src/HOL/HOLCF/Completion.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/HOLCF/Completion.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -198,7 +198,7 @@
   def b \<equiv> "\<lambda>i. LEAST j. enum j \<in> rep x \<and> \<not> enum j \<preceq> enum i"
   def c \<equiv> "\<lambda>i j. LEAST k. enum k \<in> rep x \<and> enum i \<preceq> enum k \<and> enum j \<preceq> enum k"
   def P \<equiv> "\<lambda>i. \<exists>j. enum j \<in> rep x \<and> \<not> enum j \<preceq> enum i"
-  def X \<equiv> "nat_rec a (\<lambda>n i. if P i then c i (b i) else i)"
+  def X \<equiv> "rec_nat a (\<lambda>n i. if P i then c i (b i) else i)"
   have X_0: "X 0 = a" unfolding X_def by simp
   have X_Suc: "\<And>n. X (Suc n) = (if P (X n) then c (X n) (b (X n)) else X n)"
     unfolding X_def by simp
--- a/src/HOL/HOLCF/IOA/Storage/Action.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/HOLCF/IOA/Storage/Action.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -8,9 +8,9 @@
 imports Main
 begin
 
-datatype action = New  | Loc nat | Free nat
+datatype action = New | Loc nat | Free nat
 
-lemma [cong]: "!!x. x = y ==> action_case a b c x = action_case a b c y"
+lemma [cong]: "!!x. x = y ==> case_action a b c x = case_action a b c y"
   by simp
 
 end
--- a/src/HOL/HOLCF/Library/List_Cpo.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/HOLCF/Library/List_Cpo.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -165,7 +165,7 @@
   shows "(\<Squnion>i. A i # B i) = (\<Squnion>i. A i) # (\<Squnion>i. B i)"
 by (intro lub_eqI is_lub_Cons cpo_lubI A B)
 
-lemma cont2cont_list_case:
+lemma cont2cont_case_list:
   assumes f: "cont (\<lambda>x. f x)"
   assumes g: "cont (\<lambda>x. g x)"
   assumes h1: "\<And>y ys. cont (\<lambda>x. h x y ys)"
@@ -186,17 +186,17 @@
 apply (case_tac y, simp_all add: g h1)
 done
 
-lemma cont2cont_list_case' [simp, cont2cont]:
+lemma cont2cont_case_list' [simp, cont2cont]:
   assumes f: "cont (\<lambda>x. f x)"
   assumes g: "cont (\<lambda>x. g x)"
   assumes h: "cont (\<lambda>p. h (fst p) (fst (snd p)) (snd (snd p)))"
   shows "cont (\<lambda>x. case f x of [] \<Rightarrow> g x | y # ys \<Rightarrow> h x y ys)"
-using assms by (simp add: cont2cont_list_case prod_cont_iff)
+using assms by (simp add: cont2cont_case_list prod_cont_iff)
 
 text {* The simple version (due to Joachim Breitner) is needed if the
   element type of the list is not a cpo. *}
 
-lemma cont2cont_list_case_simple [simp, cont2cont]:
+lemma cont2cont_case_list_simple [simp, cont2cont]:
   assumes "cont (\<lambda>x. f1 x)"
   assumes "\<And>y ys. cont (\<lambda>x. f2 x y ys)"
   shows "cont (\<lambda>x. case l of [] \<Rightarrow> f1 x | y # ys \<Rightarrow> f2 x y ys)"
--- a/src/HOL/HOLCF/Library/Option_Cpo.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/HOLCF/Library/Option_Cpo.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -117,7 +117,7 @@
 
 lemmas lub_Some = cont2contlubE [OF cont_Some, symmetric]
 
-lemma cont2cont_option_case:
+lemma cont2cont_case_option:
   assumes f: "cont (\<lambda>x. f x)"
   assumes g: "cont (\<lambda>x. g x)"
   assumes h1: "\<And>a. cont (\<lambda>x. h x a)"
@@ -134,16 +134,16 @@
 apply (case_tac y, simp_all add: g h1)
 done
 
-lemma cont2cont_option_case' [simp, cont2cont]:
+lemma cont2cont_case_option' [simp, cont2cont]:
   assumes f: "cont (\<lambda>x. f x)"
   assumes g: "cont (\<lambda>x. g x)"
   assumes h: "cont (\<lambda>p. h (fst p) (snd p))"
   shows "cont (\<lambda>x. case f x of None \<Rightarrow> g x | Some a \<Rightarrow> h x a)"
-using assms by (simp add: cont2cont_option_case prod_cont_iff)
+using assms by (simp add: cont2cont_case_option prod_cont_iff)
 
 text {* Simple version for when the element type is not a cpo. *}
 
-lemma cont2cont_option_case_simple [simp, cont2cont]:
+lemma cont2cont_case_option_simple [simp, cont2cont]:
   assumes "cont (\<lambda>x. f x)"
   assumes "\<And>a. cont (\<lambda>x. g x a)"
   shows "cont (\<lambda>x. case z of None \<Rightarrow> f x | Some a \<Rightarrow> g x a)"
--- a/src/HOL/HOLCF/Library/Sum_Cpo.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/HOLCF/Library/Sum_Cpo.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -145,39 +145,39 @@
 lemmas lub_Inl = cont2contlubE [OF cont_Inl, symmetric]
 lemmas lub_Inr = cont2contlubE [OF cont_Inr, symmetric]
 
-lemma cont_sum_case1:
+lemma cont_case_sum1:
   assumes f: "\<And>a. cont (\<lambda>x. f x a)"
   assumes g: "\<And>b. cont (\<lambda>x. g x b)"
   shows "cont (\<lambda>x. case y of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)"
 by (induct y, simp add: f, simp add: g)
 
-lemma cont_sum_case2: "\<lbrakk>cont f; cont g\<rbrakk> \<Longrightarrow> cont (sum_case f g)"
+lemma cont_case_sum2: "\<lbrakk>cont f; cont g\<rbrakk> \<Longrightarrow> cont (case_sum f g)"
 apply (rule contI)
 apply (erule sum_chain_cases)
 apply (simp add: cont2contlubE [OF cont_Inl, symmetric] contE)
 apply (simp add: cont2contlubE [OF cont_Inr, symmetric] contE)
 done
 
-lemma cont2cont_sum_case:
+lemma cont2cont_case_sum:
   assumes f1: "\<And>a. cont (\<lambda>x. f x a)" and f2: "\<And>x. cont (\<lambda>a. f x a)"
   assumes g1: "\<And>b. cont (\<lambda>x. g x b)" and g2: "\<And>x. cont (\<lambda>b. g x b)"
   assumes h: "cont (\<lambda>x. h x)"
   shows "cont (\<lambda>x. case h x of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)"
 apply (rule cont_apply [OF h])
-apply (rule cont_sum_case2 [OF f2 g2])
-apply (rule cont_sum_case1 [OF f1 g1])
+apply (rule cont_case_sum2 [OF f2 g2])
+apply (rule cont_case_sum1 [OF f1 g1])
 done
 
-lemma cont2cont_sum_case' [simp, cont2cont]:
+lemma cont2cont_case_sum' [simp, cont2cont]:
   assumes f: "cont (\<lambda>p. f (fst p) (snd p))"
   assumes g: "cont (\<lambda>p. g (fst p) (snd p))"
   assumes h: "cont (\<lambda>x. h x)"
   shows "cont (\<lambda>x. case h x of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)"
-using assms by (simp add: cont2cont_sum_case prod_cont_iff)
+using assms by (simp add: cont2cont_case_sum prod_cont_iff)
 
 text {* Continuity of map function. *}
 
-lemma sum_map_eq: "sum_map f g = sum_case (\<lambda>a. Inl (f a)) (\<lambda>b. Inr (g b))"
+lemma sum_map_eq: "sum_map f g = case_sum (\<lambda>a. Inl (f a)) (\<lambda>b. Inr (g b))"
 by (rule ext, case_tac x, simp_all)
 
 lemma cont2cont_sum_map [simp, cont2cont]:
--- a/src/HOL/HOLCF/Lift.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/HOLCF/Lift.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -71,23 +71,23 @@
     by (induct x) auto
 qed
 
-subsection {* Continuity of @{const lift_case} *}
+subsection {* Continuity of @{const case_lift} *}
 
-lemma lift_case_eq: "lift_case \<bottom> f x = fup\<cdot>(\<Lambda> y. f (undiscr y))\<cdot>(Rep_lift x)"
+lemma case_lift_eq: "case_lift \<bottom> f x = fup\<cdot>(\<Lambda> y. f (undiscr y))\<cdot>(Rep_lift x)"
 apply (induct x, unfold lift.cases)
 apply (simp add: Rep_lift_strict)
 apply (simp add: Def_def Abs_lift_inverse)
 done
 
-lemma cont2cont_lift_case [simp]:
-  "\<lbrakk>\<And>y. cont (\<lambda>x. f x y); cont g\<rbrakk> \<Longrightarrow> cont (\<lambda>x. lift_case \<bottom> (f x) (g x))"
-unfolding lift_case_eq by (simp add: cont_Rep_lift)
+lemma cont2cont_case_lift [simp]:
+  "\<lbrakk>\<And>y. cont (\<lambda>x. f x y); cont g\<rbrakk> \<Longrightarrow> cont (\<lambda>x. case_lift \<bottom> (f x) (g x))"
+unfolding case_lift_eq by (simp add: cont_Rep_lift)
 
 subsection {* Further operations *}
 
 definition
   flift1 :: "('a \<Rightarrow> 'b::pcpo) \<Rightarrow> ('a lift \<rightarrow> 'b)"  (binder "FLIFT " 10)  where
-  "flift1 = (\<lambda>f. (\<Lambda> x. lift_case \<bottom> f x))"
+  "flift1 = (\<lambda>f. (\<Lambda> x. case_lift \<bottom> f x))"
 
 translations
   "\<Lambda>(XCONST Def x). t" => "CONST flift1 (\<lambda>x. t)"
--- a/src/HOL/HOLCF/Product_Cpo.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/HOLCF/Product_Cpo.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -213,7 +213,7 @@
 
 lemmas cont2cont_snd [simp, cont2cont] = cont_compose [OF cont_snd]
 
-lemma cont2cont_prod_case:
+lemma cont2cont_case_prod:
   assumes f1: "\<And>a b. cont (\<lambda>x. f x a b)"
   assumes f2: "\<And>x b. cont (\<lambda>a. f x a b)"
   assumes f3: "\<And>x a. cont (\<lambda>b. f x a b)"
@@ -233,7 +233,7 @@
   shows "cont f"
 proof -
   have "cont (\<lambda>(x, y). f (x, y))"
-    by (intro cont2cont_prod_case f1 f2 cont2cont)
+    by (intro cont2cont_case_prod f1 f2 cont2cont)
   thus "cont f"
     by (simp only: split_eta)
 qed
@@ -246,11 +246,11 @@
 apply (simp only: prod_contI)
 done
 
-lemma cont2cont_prod_case' [simp, cont2cont]:
+lemma cont2cont_case_prod' [simp, cont2cont]:
   assumes f: "cont (\<lambda>p. f (fst p) (fst (snd p)) (snd (snd p)))"
   assumes g: "cont (\<lambda>x. g x)"
-  shows "cont (\<lambda>x. prod_case (f x) (g x))"
-using assms by (simp add: cont2cont_prod_case prod_cont_iff)
+  shows "cont (\<lambda>x. case_prod (f x) (g x))"
+using assms by (simp add: cont2cont_case_prod prod_cont_iff)
 
 text {* The simple version (due to Joachim Breitner) is needed if
   either element type of the pair is not a cpo. *}
@@ -262,10 +262,10 @@
 
 text {* Admissibility of predicates on product types. *}
 
-lemma adm_prod_case [simp]:
+lemma adm_case_prod [simp]:
   assumes "adm (\<lambda>x. P x (fst (f x)) (snd (f x)))"
   shows "adm (\<lambda>x. case f x of (a, b) \<Rightarrow> P x a b)"
-unfolding prod_case_beta using assms .
+unfolding case_prod_beta using assms .
 
 subsection {* Compactness and chain-finiteness *}
 
--- a/src/HOL/Hilbert_Choice.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Hilbert_Choice.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -289,7 +289,7 @@
   shows "\<exists>f. inj (f::nat \<Rightarrow> 'a) \<and> range f \<subseteq> S"
   -- {* Courtesy of Stephan Merz *}
 proof -
-  def Sseq \<equiv> "nat_rec S (\<lambda>n T. T - {SOME e. e \<in> T})"
+  def Sseq \<equiv> "rec_nat S (\<lambda>n T. T - {SOME e. e \<in> T})"
   def pick \<equiv> "\<lambda>n. (SOME e. e \<in> Sseq n)"
   { fix n have "Sseq n \<subseteq> S" "\<not> finite (Sseq n)" by (induct n) (auto simp add: Sseq_def inf) }
   moreover then have *: "\<And>n. pick n \<in> Sseq n" by (metis someI_ex pick_def ex_in_conv finite.simps)
@@ -534,8 +534,8 @@
  apply (erule exE)
  apply (erule_tac x = "{w. \<exists>i. w=f i}" in allE, blast)
 apply (erule contrapos_np, simp, clarify)
-apply (subgoal_tac "\<forall>n. nat_rec x (%i y. @z. z:Q & (z,y) :r) n \<in> Q")
- apply (rule_tac x = "nat_rec x (%i y. @z. z:Q & (z,y) :r)" in exI)
+apply (subgoal_tac "\<forall>n. rec_nat x (%i y. @z. z:Q & (z,y) :r) n \<in> Q")
+ apply (rule_tac x = "rec_nat x (%i y. @z. z:Q & (z,y) :r)" in exI)
  apply (rule allI, simp)
  apply (rule someI2_ex, blast, blast)
 apply (rule allI)
--- a/src/HOL/Hoare/hoare_syntax.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Hoare/hoare_syntax.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -28,7 +28,7 @@
 
 fun mk_abstuple [x] body = Syntax_Trans.abs_tr [x, body]
   | mk_abstuple (x :: xs) body =
-      Syntax.const @{const_syntax prod_case} $ Syntax_Trans.abs_tr [x, mk_abstuple xs body];
+      Syntax.const @{const_syntax case_prod} $ Syntax_Trans.abs_tr [x, mk_abstuple xs body];
 
 fun mk_fbody x e [y] = if eq_idt (x, y) then e else y
   | mk_fbody x e (y :: xs) =
@@ -82,21 +82,21 @@
 local
 
 fun dest_abstuple
-      (Const (@{const_syntax prod_case}, _) $ Abs (v, _, body)) =
+      (Const (@{const_syntax case_prod}, _) $ Abs (v, _, body)) =
         subst_bound (Syntax.free v, dest_abstuple body)
   | dest_abstuple (Abs (v,_, body)) = subst_bound (Syntax.free v, body)
   | dest_abstuple tm = tm;
 
-fun abs2list (Const (@{const_syntax prod_case}, _) $ Abs (x, T, t)) = Free (x, T) :: abs2list t
+fun abs2list (Const (@{const_syntax case_prod}, _) $ Abs (x, T, t)) = Free (x, T) :: abs2list t
   | abs2list (Abs (x, T, t)) = [Free (x, T)]
   | abs2list _ = [];
 
-fun mk_ts (Const (@{const_syntax prod_case}, _) $ Abs (x, _, t)) = mk_ts t
+fun mk_ts (Const (@{const_syntax case_prod}, _) $ Abs (x, _, t)) = mk_ts t
   | mk_ts (Abs (x, _, t)) = mk_ts t
   | mk_ts (Const (@{const_syntax Pair}, _) $ a $ b) = a :: mk_ts b
   | mk_ts t = [t];
 
-fun mk_vts (Const (@{const_syntax prod_case},_) $ Abs (x, _, t)) =
+fun mk_vts (Const (@{const_syntax case_prod},_) $ Abs (x, _, t)) =
       (Syntax.free x :: abs2list t, mk_ts t)
   | mk_vts (Abs (x, _, t)) = ([Syntax.free x], [t])
   | mk_vts t = raise Match;
@@ -106,7 +106,7 @@
       if t = Bound i then find_ch vts (i - 1) xs
       else (true, (v, subst_bounds (xs, t)));
 
-fun is_f (Const (@{const_syntax prod_case}, _) $ Abs (x, _, t)) = true
+fun is_f (Const (@{const_syntax case_prod}, _) $ Abs (x, _, t)) = true
   | is_f (Abs (x, _, t)) = true
   | is_f t = false;
 
--- a/src/HOL/Hoare/hoare_tac.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Hoare/hoare_tac.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -18,7 +18,7 @@
 local
 
 (** maps (%x1 ... xn. t) to [x1,...,xn] **)
-fun abs2list (Const (@{const_name prod_case}, _) $ Abs (x, T, t)) = Free (x, T) :: abs2list t
+fun abs2list (Const (@{const_name case_prod}, _) $ Abs (x, T, t)) = Free (x, T) :: abs2list t
   | abs2list (Abs (x, T, t)) = [Free (x, T)]
   | abs2list _ = [];
 
@@ -39,7 +39,7 @@
             Abs (_, T, _) => T
           | Const (_, Type (_, [_, Type (_, [T, _])])) $ _ => T);
       in
-        Const (@{const_name prod_case},
+        Const (@{const_name case_prod},
             (T --> T2 --> HOLogic.boolT) --> HOLogic.mk_prodT (T, T2) --> HOLogic.boolT) $
           absfree (x, T) z
       end;
--- a/src/HOL/Hoare_Parallel/Graph.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Hoare_Parallel/Graph.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -53,6 +53,7 @@
 apply(case_tac "list")
  apply force
 apply simp
+apply(rename_tac lista)
 apply(rotate_tac -2)
 apply(erule_tac x = "0" in all_dupE)
 apply simp
--- a/src/HOL/Hoare_Parallel/RG_Hoare.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Hoare.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -636,6 +636,7 @@
   prefer 2
   apply force
  apply(case_tac xsa,simp,simp)
+ apply(rename_tac list)
  apply(rule_tac x="(Some Pa, sa) #(Some Pa, t) # list" in exI,simp)
  apply(rule conjI,erule CptnEnv)
  apply(simp (no_asm_use) add:lift_def)
@@ -733,6 +734,7 @@
  apply(case_tac xs,simp add:cp_def)
  apply clarify
  apply (simp del:map.simps)
+ apply (rename_tac list)
  apply(subgoal_tac "(map (lift Q) ((a, b) # list))\<noteq>[]")
   apply(drule last_conv_nth)
   apply (simp del:map.simps)
@@ -1032,6 +1034,7 @@
  apply(drule last_conv_nth)
  apply (simp del:map.simps last.simps)
  apply(simp add:nth_append del:last.simps)
+ apply(rename_tac a list)
  apply(subgoal_tac "((Some (While b P), snd (last ((Some P, sa) # xs))) # a # list)\<noteq>[]")
   apply(drule last_conv_nth)
   apply (simp del:map.simps last.simps)
@@ -1349,6 +1352,7 @@
 apply(subgoal_tac "xs\<noteq>[]")
  prefer 2
  apply simp
+apply(rename_tac a list)
 apply(thin_tac "xs = a # list")
 apply(simp add:par_com_validity_def par_comm_def)
 apply clarify
--- a/src/HOL/Hoare_Parallel/RG_Tran.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Tran.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -838,6 +838,7 @@
   apply(case_tac x)
    apply(force elim:par_cptn.cases)
   apply simp
+  apply(rename_tac a list)
   apply(erule_tac x="list" in allE)
   apply clarify
   apply simp
--- a/src/HOL/Imperative_HOL/ex/Linked_Lists.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Imperative_HOL/ex/Linked_Lists.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -692,7 +692,7 @@
 
 
 
-lemma sum_distrib: "sum_case fl fr (case x of Empty \<Rightarrow> y | Node v n \<Rightarrow> (z v n)) = (case x of Empty \<Rightarrow> sum_case fl fr y | Node v n \<Rightarrow> sum_case fl fr (z v n))"
+lemma sum_distrib: "case_sum fl fr (case x of Empty \<Rightarrow> y | Node v n \<Rightarrow> (z v n)) = (case x of Empty \<Rightarrow> case_sum fl fr y | Node v n \<Rightarrow> case_sum fl fr (z v n))"
 by (cases x) auto
 
 subsection {* Induction refinement by applying the abstraction function to our induct rule *}
--- a/src/HOL/Imperative_HOL/ex/SatChecker.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Imperative_HOL/ex/SatChecker.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -459,7 +459,7 @@
                 else raise(''No empty clause''))
   }"
 
-lemma effect_option_case:
+lemma effect_case_option:
   assumes "effect (case x of None \<Rightarrow> n | Some y \<Rightarrow> s y) h h' r"
   obtains "x = None" "effect n h h' r"
          | y where "x = Some y" "effect (s y) h h' r" 
@@ -500,7 +500,7 @@
   }
   with assms show ?thesis
     unfolding res_thm2.simps get_clause_def
-    by (elim effect_bindE effect_ifE effect_nthE effect_raiseE effect_returnE effect_option_case) auto
+    by (elim effect_bindE effect_ifE effect_nthE effect_raiseE effect_returnE effect_case_option) auto
 qed
 
 lemma foldM_Inv2:
@@ -543,7 +543,7 @@
   show ?thesis
     apply auto
     apply (auto simp: get_clause_def elim!: effect_bindE effect_nthE)
-    apply (auto elim!: effect_bindE effect_nthE effect_option_case effect_raiseE
+    apply (auto elim!: effect_bindE effect_nthE effect_case_option effect_raiseE
       effect_returnE effect_updE)
     apply (frule foldM_Inv2)
     apply assumption
--- a/src/HOL/Import/HOL_Light_Maps.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Import/HOL_Light_Maps.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -210,14 +210,14 @@
 
 lemma sum_RECURSION:
   "\<forall>Inl' Inr'. \<exists>fn. (\<forall>a :: 'A. fn (Inl a) = (Inl' a :: 'Z)) \<and> (\<forall>a :: 'B. fn (Inr a) = Inr' a)"
-  by (intro allI, rule_tac x="sum_case Inl' Inr'" in exI) auto
+  by (intro allI, rule_tac x="case_sum Inl' Inr'" in exI) auto
 
-lemma OUTL[import_const "OUTL" : "Sum_Type.Projl"]:
-  "Sum_Type.Projl (Inl x) = x"
+lemma OUTL[import_const "OUTL" : "Sum_Type.projl"]:
+  "Sum_Type.projl (Inl x) = x"
   by simp
 
-lemma OUTR[import_const "OUTR" : "Sum_Type.Projr"]:
-  "Sum_Type.Projr (Inr y) = y"
+lemma OUTR[import_const "OUTR" : "Sum_Type.projr"]:
+  "Sum_Type.projr (Inr y) = y"
   by simp
 
 import_type_map list : List.list
@@ -230,13 +230,13 @@
 
 lemma list_RECURSION:
  "\<forall>nil' cons'. \<exists>fn\<Colon>'A list \<Rightarrow> 'Z. fn [] = nil' \<and> (\<forall>(a0\<Colon>'A) a1\<Colon>'A list. fn (a0 # a1) = cons' a0 a1 (fn a1))"
-  by (intro allI, rule_tac x="list_rec nil' cons'" in exI) auto
+  by (intro allI, rule_tac x="rec_list nil' cons'" in exI) auto
 
-lemma HD[import_const HD : List.hd]:
+lemma HD[import_const HD : List.list.hd]:
   "hd ((h\<Colon>'A) # t) = h"
   by simp
 
-lemma TL[import_const TL : List.tl]:
+lemma TL[import_const TL : List.list.tl]:
   "tl ((h\<Colon>'A) # t) = t"
   by simp
 
--- a/src/HOL/Induct/QuoNestedDataType.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Induct/QuoNestedDataType.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -212,6 +212,7 @@
 
 lemma ExpList_rep: "\<exists>Us. z = Abs_ExpList Us"
 apply (induct z)
+apply (rename_tac [2] a b)
 apply (rule_tac [2] z=a in eq_Abs_Exp)
 apply (auto simp add: Abs_ExpList_def Cons_eq_map_conv intro: exprel_refl)
 done
--- a/src/HOL/Int.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Int.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -6,7 +6,7 @@
 header {* The Integers as Equivalence Classes over Pairs of Natural Numbers *} 
 
 theory Int
-imports Equiv_Relations Wellfounded Power Quotient Fun_Def
+imports Equiv_Relations Power Quotient Fun_Def
 begin
 
 subsection {* Definition of integers as a quotient type *}
--- a/src/HOL/Lazy_Sequence.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Lazy_Sequence.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -35,12 +35,12 @@
   "size (xq :: 'a lazy_sequence) = 0"
   by (cases xq) simp
 
-lemma lazy_sequence_case [simp]:
-  "lazy_sequence_case f xq = f (list_of_lazy_sequence xq)"
+lemma case_lazy_sequence [simp]:
+  "case_lazy_sequence f xq = f (list_of_lazy_sequence xq)"
   by (cases xq) auto
 
-lemma lazy_sequence_rec [simp]:
-  "lazy_sequence_rec f xq = f (list_of_lazy_sequence xq)"
+lemma rec_lazy_sequence [simp]:
+  "rec_lazy_sequence f xq = f (list_of_lazy_sequence xq)"
   by (cases xq) auto
 
 definition Lazy_Sequence :: "(unit \<Rightarrow> ('a \<times> 'a lazy_sequence) option) \<Rightarrow> 'a lazy_sequence"
@@ -71,8 +71,8 @@
   "yield (Lazy_Sequence f) = f ()"
   by (cases "f ()") (simp_all add: yield_def split_def)
 
-lemma case_yield_eq [simp]: "option_case g h (yield xq) =
-  list_case g (\<lambda>x. curry h x \<circ> lazy_sequence_of_list) (list_of_lazy_sequence xq)"
+lemma case_yield_eq [simp]: "case_option g h (yield xq) =
+  case_list g (\<lambda>x. curry h x \<circ> lazy_sequence_of_list) (list_of_lazy_sequence xq)"
   by (cases "list_of_lazy_sequence xq") (simp_all add: yield_def)
 
 lemma lazy_sequence_size_code [code]:
@@ -346,4 +346,3 @@
   if_seq_def those_def not_seq_def product_def 
 
 end
-
--- a/src/HOL/Library/AList.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Library/AList.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -79,7 +79,7 @@
   by (simp add: update_conv')
 
 definition updates :: "'key list \<Rightarrow> 'val list \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list" where
-  "updates ks vs = fold (prod_case update) (zip ks vs)"
+  "updates ks vs = fold (case_prod update) (zip ks vs)"
 
 lemma updates_simps [simp]:
   "updates [] vs ps = ps"
@@ -94,7 +94,7 @@
 
 lemma updates_conv': "map_of (updates ks vs al) = (map_of al)(ks[\<mapsto>]vs)"
 proof -
-  have "map_of \<circ> fold (prod_case update) (zip ks vs) =
+  have "map_of \<circ> fold (case_prod update) (zip ks vs) =
     fold (\<lambda>(k, v) f. f(k \<mapsto> v)) (zip ks vs) \<circ> map_of"
     by (rule fold_commute) (auto simp add: fun_eq_iff update_conv')
   then show ?thesis by (auto simp add: updates_def fun_eq_iff map_upds_fold_map_upd foldl_conv_fold split_def)
@@ -111,9 +111,9 @@
        (\<lambda>(k, v) al. if k \<in> set al then al else al @ [k])
        (zip ks vs) (map fst al))"
     by (rule fold_invariant [of "zip ks vs" "\<lambda>_. True"]) (auto intro: assms)
-  moreover have "map fst \<circ> fold (prod_case update) (zip ks vs) =
+  moreover have "map fst \<circ> fold (case_prod update) (zip ks vs) =
     fold (\<lambda>(k, v) al. if k \<in> set al then al else al @ [k]) (zip ks vs) \<circ> map fst"
-    by (rule fold_commute) (simp add: update_keys split_def prod_case_beta comp_def)
+    by (rule fold_commute) (simp add: update_keys split_def case_prod_beta comp_def)
   ultimately show ?thesis by (simp add: updates_def fun_eq_iff)
 qed
 
@@ -339,9 +339,9 @@
 lemma clearjunk_updates:
   "clearjunk (updates ks vs al) = updates ks vs (clearjunk al)"
 proof -
-  have "clearjunk \<circ> fold (prod_case update) (zip ks vs) =
-    fold (prod_case update) (zip ks vs) \<circ> clearjunk"
-    by (rule fold_commute) (simp add: clearjunk_update prod_case_beta o_def)
+  have "clearjunk \<circ> fold (case_prod update) (zip ks vs) =
+    fold (case_prod update) (zip ks vs) \<circ> clearjunk"
+    by (rule fold_commute) (simp add: clearjunk_update case_prod_beta o_def)
   then show ?thesis by (simp add: updates_def fun_eq_iff)
 qed
 
@@ -444,9 +444,9 @@
 lemma merge_conv':
   "map_of (merge xs ys) = map_of xs ++ map_of ys"
 proof -
-  have "map_of \<circ> fold (prod_case update) (rev ys) =
+  have "map_of \<circ> fold (case_prod update) (rev ys) =
     fold (\<lambda>(k, v) m. m(k \<mapsto> v)) (rev ys) \<circ> map_of"
-    by (rule fold_commute) (simp add: update_conv' prod_case_beta split_def fun_eq_iff)
+    by (rule fold_commute) (simp add: update_conv' case_prod_beta split_def fun_eq_iff)
   then show ?thesis
     by (simp add: merge_def map_add_map_of_foldr foldr_conv_fold fun_eq_iff)
 qed
--- a/src/HOL/Library/Bit.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Library/Bit.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -102,10 +102,10 @@
 begin
 
 definition plus_bit_def:
-  "x + y = bit_case y (bit_case 1 0 y) x"
+  "x + y = case_bit y (case_bit 1 0 y) x"
 
 definition times_bit_def:
-  "x * y = bit_case 0 y x"
+  "x * y = case_bit 0 y x"
 
 definition uminus_bit_def [simp]:
   "- x = (x :: bit)"
@@ -167,7 +167,7 @@
 
 definition of_bit :: "bit \<Rightarrow> 'a"
 where
-  "of_bit b = bit_case 0 1 b" 
+  "of_bit b = case_bit 0 1 b" 
 
 lemma of_bit_eq [simp, code]:
   "of_bit 0 = 0"
--- a/src/HOL/Library/Code_Abstract_Nat.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Library/Code_Abstract_Nat.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -24,7 +24,7 @@
 *}
 
 lemma [code, code_unfold]:
-  "nat_case = (\<lambda>f g n. if n = 0 then f else g (n - 1))"
+  "case_nat = (\<lambda>f g n. if n = 0 then f else g (n - 1))"
   by (auto simp add: fun_eq_iff dest!: gr0_implies_Suc)
 
 
--- a/src/HOL/Library/Countable.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Library/Countable.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -90,7 +90,7 @@
 text {* Options *}
 
 instance option :: (countable) countable
-  by (rule countable_classI [of "option_case 0 (Suc \<circ> to_nat)"])
+  by (rule countable_classI [of "case_option 0 (Suc \<circ> to_nat)"])
     (simp split: option.split_asm)
 
 
--- a/src/HOL/Library/Countable_Set_Type.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Library/Countable_Set_Type.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -140,7 +140,7 @@
   have L: "countable ?L'" by auto
   hence *: "rcset R' = ?L'" unfolding R'_def by (intro rcset_to_rcset)
   thus ?R unfolding Grp_def relcompp.simps conversep.simps
-  proof (intro CollectI prod_caseI exI[of _ a] exI[of _ b] exI[of _ R'] conjI refl)
+  proof (intro CollectI case_prodI exI[of _ a] exI[of _ b] exI[of _ R'] conjI refl)
     from * `?L` show "a = cimage fst R'" by transfer (auto simp: image_def Collect_Int_Times)
   next
     from * `?L` show "b = cimage snd R'" by transfer (auto simp: image_def Collect_Int_Times)
--- a/src/HOL/Library/FSet.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Library/FSet.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -1017,7 +1017,7 @@
   have "finite ?L'" by (intro finite_Int[OF disjI2] finite_cartesian_product) (transfer, simp)+
   hence *: "fset R' = ?L'" unfolding R'_def by (intro fset_to_fset)
   show ?R unfolding Grp_def relcompp.simps conversep.simps
-  proof (intro CollectI prod_caseI exI[of _ a] exI[of _ b] exI[of _ R'] conjI refl)
+  proof (intro CollectI case_prodI exI[of _ a] exI[of _ b] exI[of _ R'] conjI refl)
     from * show "a = fimage fst R'" using conjunct1[OF `?L`]
       by (transfer, auto simp add: image_def Int_def split: prod.splits)
     from * show "b = fimage snd R'" using conjunct2[OF `?L`]
--- a/src/HOL/Library/IArray.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Library/IArray.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -63,11 +63,11 @@
 by (cases as) simp
 
 lemma [code]:
-"iarray_rec f as = f (IArray.list_of as)"
+"rec_iarray f as = f (IArray.list_of as)"
 by (cases as) simp
 
 lemma [code]:
-"iarray_case f as = f (IArray.list_of as)"
+"case_iarray f as = f (IArray.list_of as)"
 by (cases as) simp
 
 lemma [code]:
--- a/src/HOL/Library/Multiset.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Library/Multiset.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -955,6 +955,7 @@
 lemma distinct_count_atmost_1:
   "distinct x = (! a. count (multiset_of x) a = (if a \<in> set x then 1 else 0))"
 apply (induct x, simp, rule iffI, simp_all)
+apply (rename_tac a b)
 apply (rule conjI)
 apply (simp_all add: set_of_multiset_of [THEN sym] del: set_of_multiset_of)
 apply (erule_tac x = a in allE, simp, clarify)
--- a/src/HOL/Library/Polynomial.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Library/Polynomial.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -145,9 +145,9 @@
   with that show thesis .
 qed
 
-lemma almost_everywhere_zero_nat_case:
+lemma almost_everywhere_zero_case_nat:
   assumes "almost_everywhere_zero f"
-  shows "almost_everywhere_zero (nat_case a f)"
+  shows "almost_everywhere_zero (case_nat a f)"
   using assms
   by (auto intro!: almost_everywhere_zeroI elim!: almost_everywhere_zeroE split: nat.split)
     blast
@@ -258,8 +258,8 @@
 subsection {* List-style constructor for polynomials *}
 
 lift_definition pCons :: "'a::zero \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
-  is "\<lambda>a p. nat_case a (coeff p)"
-  using coeff_almost_everywhere_zero by (rule almost_everywhere_zero_nat_case)
+  is "\<lambda>a p. case_nat a (coeff p)"
+  using coeff_almost_everywhere_zero by (rule almost_everywhere_zero_case_nat)
 
 lemmas coeff_pCons = pCons.rep_eq
 
@@ -405,8 +405,8 @@
 proof -
   { fix ms :: "nat list" and f :: "nat \<Rightarrow> 'a" and x :: "'a"
     assume "\<forall>m\<in>set ms. m > 0"
-    then have "map (nat_case x f) ms = map f (map (\<lambda>n. n - 1) ms)"
-      by (induct ms) (auto, metis Suc_pred' nat_case_Suc) }
+    then have "map (case_nat x f) ms = map f (map (\<lambda>n. n - 1) ms)"
+      by (induct ms) (auto, metis Suc_pred' nat.cases(2)) }
   note * = this
   show ?thesis
     by (simp add: coeffs_def * upt_conv_Cons coeff_pCons map_decr_upt One_nat_def del: upt_Suc)
@@ -452,7 +452,7 @@
 lemma coeff_Poly_eq:
   "coeff (Poly xs) n = nth_default 0 xs n"
   apply (induct xs arbitrary: n) apply simp_all
-  by (metis nat_case_0 nat_case_Suc not0_implies_Suc nth_default_Cons_0 nth_default_Cons_Suc pCons.rep_eq)
+  by (metis nat.cases not0_implies_Suc nth_default_Cons_0 nth_default_Cons_Suc pCons.rep_eq)
 
 lemma nth_default_coeffs_eq:
   "nth_default 0 (coeffs p) = coeff p"
--- a/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -22,7 +22,7 @@
 
 section {* Pairs *}
 
-setup {* Predicate_Compile_Data.ignore_consts [@{const_name fst}, @{const_name snd}, @{const_name prod_case}] *}
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name fst}, @{const_name snd}, @{const_name case_prod}] *}
 
 section {* Bounded quantifiers *}
 
--- a/src/HOL/Library/Quotient_Product.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Library/Quotient_Product.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -78,7 +78,7 @@
 
 lemma split_rsp [quot_respect]:
   shows "((R1 ===> R2 ===> (op =)) ===> (prod_rel R1 R2) ===> (op =)) split split"
-  by (rule prod_case_transfer)
+  by (rule case_prod_transfer)
 
 lemma split_prs [quot_preserve]:
   assumes q1: "Quotient3 R1 Abs1 Rep1"
--- a/src/HOL/Library/RBT.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Library/RBT.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -135,7 +135,7 @@
   by transfer (rule rbt_lookup_map)
 
 lemma fold_fold:
-  "fold f t = List.fold (prod_case f) (entries t)"
+  "fold f t = List.fold (case_prod f) (entries t)"
   by transfer (rule RBT_Impl.fold_def)
 
 lemma impl_of_empty:
@@ -175,7 +175,7 @@
   by transfer (simp add: keys_entries)
 
 lemma fold_def_alt:
-  "fold f t = List.fold (prod_case f) (entries t)"
+  "fold f t = List.fold (case_prod f) (entries t)"
   by transfer (auto simp: RBT_Impl.fold_def)
 
 lemma distinct_entries: "distinct (List.map fst (entries t))"
--- a/src/HOL/Library/RBT_Impl.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Library/RBT_Impl.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -1067,7 +1067,7 @@
 subsection {* Folding over entries *}
 
 definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c" where
-  "fold f t = List.fold (prod_case f) (entries t)"
+  "fold f t = List.fold (case_prod f) (entries t)"
 
 lemma fold_simps [simp]:
   "fold f Empty = id"
@@ -1110,10 +1110,10 @@
 proof -
   obtain ys where "ys = rev xs" by simp
   have "\<And>t. is_rbt t \<Longrightarrow>
-    rbt_lookup (List.fold (prod_case rbt_insert) ys t) = rbt_lookup t ++ map_of (rev ys)"
-      by (induct ys) (simp_all add: rbt_bulkload_def rbt_lookup_rbt_insert prod_case_beta)
+    rbt_lookup (List.fold (case_prod rbt_insert) ys t) = rbt_lookup t ++ map_of (rev ys)"
+      by (induct ys) (simp_all add: rbt_bulkload_def rbt_lookup_rbt_insert case_prod_beta)
   from this Empty_is_rbt have
-    "rbt_lookup (List.fold (prod_case rbt_insert) (rev xs) Empty) = rbt_lookup Empty ++ map_of xs"
+    "rbt_lookup (List.fold (case_prod rbt_insert) (rev xs) Empty) = rbt_lookup Empty ++ map_of xs"
      by (simp add: `ys = rev xs`)
   then show ?thesis by (simp add: rbt_bulkload_def rbt_lookup_Empty foldr_conv_fold)
 qed
@@ -1167,7 +1167,7 @@
           apfst (Branch B t1 k v) (rbtreeify_g n' kvs')
       else case rbtreeify_f n' kvs of (t1, (k, v) # kvs') \<Rightarrow>
           apfst (Branch B t1 k v) (rbtreeify_f n' kvs'))"
-by(subst rbtreeify_f.simps)(simp only: Let_def divmod_nat_div_mod prod.simps)
+by (subst rbtreeify_f.simps) (simp only: Let_def divmod_nat_div_mod prod.case)
 
 lemma rbtreeify_g_code [code]:
   "rbtreeify_g n kvs =
@@ -1178,7 +1178,7 @@
           apfst (Branch B t1 k v) (rbtreeify_g n' kvs')
       else case rbtreeify_f n' kvs of (t1, (k, v) # kvs') \<Rightarrow>
           apfst (Branch B t1 k v) (rbtreeify_g n' kvs'))"
-by(subst rbtreeify_g.simps)(simp only: Let_def divmod_nat_div_mod prod.simps)
+by(subst rbtreeify_g.simps)(simp only: Let_def divmod_nat_div_mod prod.case)
 
 lemma Suc_double_half: "Suc (2 * n) div 2 = n"
 by simp
@@ -1250,8 +1250,8 @@
       with "1.prems" False obtain t1 k' v' kvs''
         where kvs'': "rbtreeify_f (n div 2) kvs = (t1, (k', v') # kvs'')"
          by(cases ?rest1)(auto simp add: snd_def split: prod.split_asm)
-      note this also note prod.simps(2) also note list.simps(5) 
-      also note prod.simps(2) also note snd_apfst
+      note this also note prod.case also note list.simps(5) 
+      also note prod.case also note snd_apfst
       also have "0 < n div 2" "n div 2 \<le> Suc (length kvs'')" 
         using len "1.prems" False unfolding kvs'' by simp_all
       with True kvs''[symmetric] refl refl
@@ -1276,8 +1276,8 @@
       with "1.prems" `\<not> n \<le> 1` obtain t1 k' v' kvs''
         where kvs'': "rbtreeify_f (n div 2) kvs = (t1, (k', v') # kvs'')"
         by(cases ?rest1)(auto simp add: snd_def split: prod.split_asm)
-      note this also note prod.simps(2) also note list.simps(5) 
-      also note prod.simps(2) also note snd_apfst
+      note this also note prod.case also note list.simps(5)
+      also note prod.case also note snd_apfst
       also have "n div 2 \<le> length kvs''" 
         using len "1.prems" False unfolding kvs'' by simp arith
       with False kvs''[symmetric] refl refl
@@ -1315,8 +1315,8 @@
       with "2.prems" obtain t1 k' v' kvs''
         where kvs'': "rbtreeify_g (n div 2) kvs = (t1, (k', v') # kvs'')"
         by(cases ?rest1)(auto simp add: snd_def split: prod.split_asm)
-      note this also note prod.simps(2) also note list.simps(5) 
-      also note prod.simps(2) also note snd_apfst
+      note this also note prod.case also note list.simps(5) 
+      also note prod.case also note snd_apfst
       also have "n div 2 \<le> Suc (length kvs'')" 
         using len "2.prems" unfolding kvs'' by simp
       with True kvs''[symmetric] refl refl `0 < n div 2`
@@ -1341,8 +1341,8 @@
       with "2.prems" `1 < n` False obtain t1 k' v' kvs'' 
         where kvs'': "rbtreeify_f (n div 2) kvs = (t1, (k', v') # kvs'')"
         by(cases ?rest1)(auto simp add: snd_def split: prod.split_asm, arith)
-      note this also note prod.simps(2) also note list.simps(5) 
-      also note prod.simps(2) also note snd_apfst
+      note this also note prod.case also note list.simps(5) 
+      also note prod.case also note snd_apfst
       also have "n div 2 \<le> Suc (length kvs'')" 
         using len "2.prems" False unfolding kvs'' by simp arith
       with False kvs''[symmetric] refl refl `0 < n div 2`
@@ -1748,14 +1748,14 @@
 
 hide_type (open) compare
 hide_const (open)
-  compare_height skip_black skip_red LT GT EQ compare_case compare_rec 
-  Abs_compare Rep_compare compare_rep_set
+  compare_height skip_black skip_red LT GT EQ case_compare rec_compare
+  Abs_compare Rep_compare rep_set_compare
 hide_fact (open)
   Abs_compare_cases Abs_compare_induct Abs_compare_inject Abs_compare_inverse
   Rep_compare Rep_compare_cases Rep_compare_induct Rep_compare_inject Rep_compare_inverse
   compare.simps compare.exhaust compare.induct compare.recs compare.simps
   compare.size compare.case_cong compare.weak_case_cong compare.cases 
-  compare.nchotomy compare.split compare.split_asm compare_rec_def
+  compare.nchotomy compare.split compare.split_asm rec_compare_def
   compare.eq.refl compare.eq.simps
   compare.EQ_def compare.GT_def compare.LT_def
   equal_compare_def
--- a/src/HOL/Library/Transitive_Closure_Table.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Library/Transitive_Closure_Table.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -1,6 +1,6 @@
 (* Author: Stefan Berghofer, Lukas Bulwahn, TU Muenchen *)
 
-header {* A tabled implementation of the reflexive transitive closure *}
+header {* A table-based implementation of the reflexive transitive closure *}
 
 theory Transitive_Closure_Table
 imports Main
--- a/src/HOL/Library/refute.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Library/refute.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -3163,8 +3163,8 @@
    add_interpreter "lfp" lfp_interpreter #>
    add_interpreter "gfp" gfp_interpreter #>
 *)
-   add_interpreter "Product_Type.fst" Product_Type_fst_interpreter #>
-   add_interpreter "Product_Type.snd" Product_Type_snd_interpreter #>
+   add_interpreter "Product_Type.prod.fst" Product_Type_fst_interpreter #>
+   add_interpreter "Product_Type.prod.snd" Product_Type_snd_interpreter #>
    add_printer "stlc" stlc_printer #>
    add_printer "set" set_printer #>
    add_printer "IDT"  IDT_printer;
--- a/src/HOL/Lifting_Option.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Lifting_Option.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -26,7 +26,7 @@
   unfolding option_rel_def by simp_all
 
 abbreviation (input) option_pred :: "('a \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> bool" where
-  "option_pred \<equiv> option_case True"
+  "option_pred \<equiv> case_option True"
 
 lemma option_rel_eq [relator_eq]:
   "option_rel (op =) = (op =)"
@@ -100,8 +100,8 @@
 lemma Some_transfer [transfer_rule]: "(A ===> option_rel A) Some Some"
   unfolding fun_rel_def by simp
 
-lemma option_case_transfer [transfer_rule]:
-  "(B ===> (A ===> B) ===> option_rel A ===> B) option_case option_case"
+lemma case_option_transfer [transfer_rule]:
+  "(B ===> (A ===> B) ===> option_rel A ===> B) case_option case_option"
   unfolding fun_rel_def split_option_all by simp
 
 lemma option_map_transfer [transfer_rule]:
@@ -115,57 +115,4 @@
 
 end
 
-
-subsubsection {* BNF setup *}
-
-lemma option_rec_conv_option_case: "option_rec = option_case"
-by (simp add: fun_eq_iff split: option.split)
-
-bnf "'a option"
-  map: Option.map
-  sets: Option.set
-  bd: natLeq
-  wits: None
-  rel: option_rel
-proof -
-  show "Option.map id = id" by (rule Option.map.id)
-next
-  fix f g
-  show "Option.map (g \<circ> f) = Option.map g \<circ> Option.map f"
-    by (auto simp add: fun_eq_iff Option.map_def split: option.split)
-next
-  fix f g x
-  assume "\<And>z. z \<in> Option.set x \<Longrightarrow> f z = g z"
-  thus "Option.map f x = Option.map g x"
-    by (simp cong: Option.map_cong)
-next
-  fix f
-  show "Option.set \<circ> Option.map f = op ` f \<circ> Option.set"
-    by fastforce
-next
-  show "card_order natLeq" by (rule natLeq_card_order)
-next
-  show "cinfinite natLeq" by (rule natLeq_cinfinite)
-next
-  fix x
-  show "|Option.set x| \<le>o natLeq"
-    by (cases x) (simp_all add: ordLess_imp_ordLeq finite_iff_ordLess_natLeq[symmetric])
-next
-  fix R S
-  show "option_rel R OO option_rel S \<le> option_rel (R OO S)"
-    by (auto simp: option_rel_def split: option.splits)
-next
-  fix z
-  assume "z \<in> Option.set None"
-  thus False by simp
-next
-  fix R
-  show "option_rel R =
-        (Grp {x. Option.set x \<subseteq> Collect (split R)} (Option.map fst))\<inverse>\<inverse> OO
-         Grp {x. Option.set x \<subseteq> Collect (split R)} (Option.map snd)"
-  unfolding option_rel_def Grp_def relcompp.simps conversep.simps fun_eq_iff prod.cases
-  by (auto simp: trans[OF eq_commute option_map_is_None] trans[OF eq_commute option_map_eq_Some]
-           split: option.splits)
-qed
-
 end
--- a/src/HOL/Lifting_Product.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Lifting_Product.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -95,8 +95,8 @@
 lemma snd_transfer [transfer_rule]: "(prod_rel A B ===> B) snd snd"
   unfolding fun_rel_def prod_rel_def by simp
 
-lemma prod_case_transfer [transfer_rule]:
-  "((A ===> B ===> C) ===> prod_rel A B ===> C) prod_case prod_case"
+lemma case_prod_transfer [transfer_rule]:
+  "((A ===> B ===> C) ===> prod_rel A B ===> C) case_prod case_prod"
   unfolding fun_rel_def prod_rel_def by simp
 
 lemma curry_transfer [transfer_rule]:
--- a/src/HOL/Lifting_Sum.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Lifting_Sum.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -10,7 +10,7 @@
 
 subsection {* Relator and predicator properties *}
 
-abbreviation (input) "sum_pred \<equiv> sum_case"
+abbreviation (input) "sum_pred \<equiv> case_sum"
 
 lemmas sum_rel_eq[relator_eq] = sum.rel_eq
 lemmas sum_rel_mono[relator_mono] = sum.rel_mono
@@ -80,8 +80,8 @@
 lemma Inr_transfer [transfer_rule]: "(B ===> sum_rel A B) Inr Inr"
   unfolding fun_rel_def by simp
 
-lemma sum_case_transfer [transfer_rule]:
-  "((A ===> C) ===> (B ===> C) ===> sum_rel A B ===> C) sum_case sum_case"
+lemma case_sum_transfer [transfer_rule]:
+  "((A ===> C) ===> (B ===> C) ===> sum_rel A B ===> C) case_sum case_sum"
   unfolding fun_rel_def sum_rel_def by (simp split: sum.split)
 
 end
--- a/src/HOL/Limits.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Limits.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -1750,7 +1750,7 @@
   assumes local: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<exists>d>0. \<forall>a b. a \<le> x \<and> x \<le> b \<and> b - a < d \<longrightarrow> P a b"
   shows "P a b"
 proof -
-  def bisect \<equiv> "nat_rec (a, b) (\<lambda>n (x, y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2))"
+  def bisect \<equiv> "rec_nat (a, b) (\<lambda>n (x, y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2))"
   def l \<equiv> "\<lambda>n. fst (bisect n)" and u \<equiv> "\<lambda>n. snd (bisect n)"
   have l[simp]: "l 0 = a" "\<And>n. l (Suc n) = (if P (l n) ((l n + u n) / 2) then (l n + u n) / 2 else l n)"
     and u[simp]: "u 0 = b" "\<And>n. u (Suc n) = (if P (l n) ((l n + u n) / 2) then u n else (l n + u n) / 2)"
--- a/src/HOL/List.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/List.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -8,9 +8,32 @@
 imports Presburger Code_Numeral Quotient Lifting_Set Lifting_Option Lifting_Product
 begin
 
-datatype 'a list =
-    Nil    ("[]")
-  | Cons 'a  "'a list"    (infixr "#" 65)
+datatype_new 'a list =
+    =: Nil (defaults tl: "[]")  ("[]")
+  | Cons (hd: 'a) (tl: "'a list")  (infixr "#" 65)
+
+datatype_new_compat list
+
+thm list.exhaust[no_vars]
+
+lemma [case_names Nil Cons, cases type: list]:
+  -- {* for backward compatibility -- names of variables differ *}
+  "(y = [] \<Longrightarrow> P) \<Longrightarrow> (\<And>a list. y = a # list \<Longrightarrow> P) \<Longrightarrow> P"
+by (rule list.exhaust)
+
+lemma [case_names Nil Cons, induct type: list]:
+  -- {* for backward compatibility -- names of variables differ *}
+  "P [] \<Longrightarrow> (\<And>a list. P list \<Longrightarrow> P (a # list)) \<Longrightarrow> P list"
+by (rule list.induct)
+
+-- {* Compatibility *}
+setup {* Sign.mandatory_path "list" *}
+
+lemmas inducts = list.induct
+lemmas recs = list.rec
+lemmas cases = list.case
+
+setup {* Sign.parent_path *}
 
 syntax
   -- {* list Enumeration *}
@@ -23,13 +46,6 @@
 
 subsection {* Basic list processing functions *}
 
-primrec hd :: "'a list \<Rightarrow> 'a" where
-"hd (x # xs) = x"
-
-primrec tl :: "'a list \<Rightarrow> 'a list" where
-"tl [] = []" |
-"tl (x # xs) = xs"
-
 primrec last :: "'a list \<Rightarrow> 'a" where
 "last (x # xs) = (if xs = [] then x else last xs)"
 
@@ -1632,10 +1648,10 @@
 lemma set_conv_nth: "set xs = {xs!i | i. i < length xs}"
 apply (induct xs, simp, simp)
 apply safe
-apply (metis nat_case_0 nth.simps zero_less_Suc)
+apply (metis nat.cases(1) nth.simps zero_less_Suc)
 apply (metis less_Suc_eq_0_disj nth_Cons_Suc)
 apply (case_tac i, simp)
-apply (metis diff_Suc_Suc nat_case_Suc nth.simps zero_less_diff)
+apply (metis diff_Suc_Suc nat.cases(2) nth.simps zero_less_diff)
 done
 
 lemma in_set_conv_nth: "(x \<in> set xs) = (\<exists>i < length xs. xs!i = x)"
@@ -3379,7 +3395,8 @@
 
 lemma distinct_length_2_or_more:
 "distinct (a # b # xs) \<longleftrightarrow> (a \<noteq> b \<and> distinct (a # xs) \<and> distinct (b # xs))"
-by (metis distinct.simps(2) hd.simps hd_in_set list.simps(2) set_ConsD set_rev_mp set_subset_Cons)
+by (metis distinct.simps(2) list.sel(1) hd_in_set list.simps(2) set_ConsD set_rev_mp
+      set_subset_Cons)
 
 lemma remdups_adj_Cons: "remdups_adj (x # xs) =
   (case remdups_adj xs of [] \<Rightarrow> [x] | y # xs \<Rightarrow> if x = y then y # xs else x # y # xs)"
@@ -4275,12 +4292,12 @@
 by pat_completeness auto
 
 lemma transpose_aux_filter_head:
-  "concat (map (list_case [] (\<lambda>h t. [h])) xss) =
+  "concat (map (case_list [] (\<lambda>h t. [h])) xss) =
   map (\<lambda>xs. hd xs) [ys\<leftarrow>xss . ys \<noteq> []]"
   by (induct xss) (auto split: list.split)
 
 lemma transpose_aux_filter_tail:
-  "concat (map (list_case [] (\<lambda>h t. [t])) xss) =
+  "concat (map (case_list [] (\<lambda>h t. [t])) xss) =
   map (\<lambda>xs. tl xs) [ys\<leftarrow>xss . ys \<noteq> []]"
   by (induct xss) (auto split: list.split)
 
@@ -4354,13 +4371,13 @@
       by (cases x) simp_all
     } note *** = this
 
-    have j_less: "j < length (transpose (xs # concat (map (list_case [] (\<lambda>h t. [t])) xss)))"
+    have j_less: "j < length (transpose (xs # concat (map (case_list [] (\<lambda>h t. [t])) xss)))"
       using "3.prems" by (simp add: transpose_aux_filter_tail length_transpose Suc)
 
     show ?thesis
       unfolding transpose.simps `i = Suc j` nth_Cons_Suc "3.hyps"[OF j_less]
       apply (auto simp: transpose_aux_filter_tail filter_map comp_def length_transpose * ** *** XS_def[symmetric])
-      apply (rule_tac y=x in list.exhaust)
+      apply (rule list.exhaust)
       by auto
   qed
 qed simp_all
@@ -6672,19 +6689,19 @@
   "(A ===> list_all2 A ===> list_all2 A) Cons Cons"
   unfolding fun_rel_def by simp
 
-lemma list_case_transfer [transfer_rule]:
+lemma case_list_transfer [transfer_rule]:
   "(B ===> (A ===> list_all2 A ===> B) ===> list_all2 A ===> B)
-    list_case list_case"
+    case_list case_list"
   unfolding fun_rel_def by (simp split: list.split)
 
-lemma list_rec_transfer [transfer_rule]:
+lemma rec_list_transfer [transfer_rule]:
   "(B ===> (A ===> list_all2 A ===> B ===> B) ===> list_all2 A ===> B)
-    list_rec list_rec"
+    rec_list rec_list"
   unfolding fun_rel_def by (clarify, erule list_all2_induct, simp_all)
 
 lemma tl_transfer [transfer_rule]:
   "(list_all2 A ===> list_all2 A) tl tl"
-  unfolding tl_def by transfer_prover
+  unfolding tl_def[abs_def] by transfer_prover
 
 lemma butlast_transfer [transfer_rule]:
   "(list_all2 A ===> list_all2 A) butlast butlast"
@@ -6876,46 +6893,4 @@
 
 end
 
-
-subsection {* BNF setup *}
-
-bnf "'a list"
-  map: map
-  sets: set
-  bd: natLeq
-  wits: Nil
-  rel: list_all2
-proof -
-  show "map id = id" by (rule List.map.id)
-next
-  fix f g
-  show "map (g o f) = map g o map f" by (rule List.map.comp[symmetric])
-next
-  fix x f g
-  assume "\<And>z. z \<in> set x \<Longrightarrow> f z = g z"
-  thus "map f x = map g x" by simp
-next
-  fix f
-  show "set o map f = image f o set" by (rule ext, unfold comp_apply, rule set_map)
-next
-  show "card_order natLeq" by (rule natLeq_card_order)
-next
-  show "cinfinite natLeq" by (rule natLeq_cinfinite)
-next
-  fix x
-  show "|set x| \<le>o natLeq"
-    by (metis List.finite_set finite_iff_ordLess_natLeq ordLess_imp_ordLeq)
-next
-  fix R S
-  show "list_all2 R OO list_all2 S \<le> list_all2 (R OO S)"
-    by (metis list_all2_OO order_refl)
-next
-  fix R
-  show "list_all2 R =
-         (Grp {x. set x \<subseteq> {(x, y). R x y}} (map fst))\<inverse>\<inverse> OO
-         Grp {x. set x \<subseteq> {(x, y). R x y}} (map snd)"
-    unfolding list_all2_def[abs_def] Grp_def fun_eq_iff relcompp.simps conversep.simps
-    by (force simp: zip_map_fst_snd)
-qed simp
-
 end
--- a/src/HOL/Matrix_LP/ComputeHOL.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Matrix_LP/ComputeHOL.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -50,9 +50,9 @@
 lemma compute_snd: "snd (x,y) = y" by simp
 lemma compute_pair_eq: "((a, b) = (c, d)) = (a = c \<and> b = d)" by auto
 
-lemma prod_case_simp: "prod_case f (x,y) = f x y" by simp
+lemma case_prod_simp: "case_prod f (x,y) = f x y" by simp
 
-lemmas compute_pair = compute_fst compute_snd compute_pair_eq prod_case_simp
+lemmas compute_pair = compute_fst compute_snd compute_pair_eq case_prod_simp
 
 (*** compute_option ***)
 
@@ -62,25 +62,25 @@
 lemma compute_None_None_eq: "(None = None) = True" by auto
 lemma compute_Some_Some_eq: "(Some x = Some y) = (x = y)" by auto
 
-definition option_case_compute :: "'b option \<Rightarrow> 'a \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a"
-  where "option_case_compute opt a f = option_case a f opt"
+definition case_option_compute :: "'b option \<Rightarrow> 'a \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a"
+  where "case_option_compute opt a f = case_option a f opt"
 
-lemma option_case_compute: "option_case = (\<lambda> a f opt. option_case_compute opt a f)"
-  by (simp add: option_case_compute_def)
+lemma case_option_compute: "case_option = (\<lambda> a f opt. case_option_compute opt a f)"
+  by (simp add: case_option_compute_def)
 
-lemma option_case_compute_None: "option_case_compute None = (\<lambda> a f. a)"
+lemma case_option_compute_None: "case_option_compute None = (\<lambda> a f. a)"
   apply (rule ext)+
-  apply (simp add: option_case_compute_def)
+  apply (simp add: case_option_compute_def)
   done
 
-lemma option_case_compute_Some: "option_case_compute (Some x) = (\<lambda> a f. f x)"
+lemma case_option_compute_Some: "case_option_compute (Some x) = (\<lambda> a f. f x)"
   apply (rule ext)+
-  apply (simp add: option_case_compute_def)
+  apply (simp add: case_option_compute_def)
   done
 
-lemmas compute_option_case = option_case_compute option_case_compute_None option_case_compute_Some
+lemmas compute_case_option = case_option_compute case_option_compute_None case_option_compute_Some
 
-lemmas compute_option = compute_the compute_None_Some_eq compute_Some_None_eq compute_None_None_eq compute_Some_Some_eq compute_option_case
+lemmas compute_option = compute_the compute_None_Some_eq compute_Some_None_eq compute_None_None_eq compute_Some_Some_eq compute_case_option
 
 (**** compute_list_length ****)
 
@@ -92,27 +92,27 @@
 
 lemmas compute_list_length = length_nil length_cons
 
-(*** compute_list_case ***)
+(*** compute_case_list ***)
 
-definition list_case_compute :: "'b list \<Rightarrow> 'a \<Rightarrow> ('b \<Rightarrow> 'b list \<Rightarrow> 'a) \<Rightarrow> 'a"
-  where "list_case_compute l a f = list_case a f l"
+definition case_list_compute :: "'b list \<Rightarrow> 'a \<Rightarrow> ('b \<Rightarrow> 'b list \<Rightarrow> 'a) \<Rightarrow> 'a"
+  where "case_list_compute l a f = case_list a f l"
 
-lemma list_case_compute: "list_case = (\<lambda> (a::'a) f (l::'b list). list_case_compute l a f)"
+lemma case_list_compute: "case_list = (\<lambda> (a::'a) f (l::'b list). case_list_compute l a f)"
   apply (rule ext)+
-  apply (simp add: list_case_compute_def)
+  apply (simp add: case_list_compute_def)
   done
 
-lemma list_case_compute_empty: "list_case_compute ([]::'b list) = (\<lambda> (a::'a) f. a)"
+lemma case_list_compute_empty: "case_list_compute ([]::'b list) = (\<lambda> (a::'a) f. a)"
   apply (rule ext)+
-  apply (simp add: list_case_compute_def)
+  apply (simp add: case_list_compute_def)
   done
 
-lemma list_case_compute_cons: "list_case_compute (u#v) = (\<lambda> (a::'a) f. (f (u::'b) v))"
+lemma case_list_compute_cons: "case_list_compute (u#v) = (\<lambda> (a::'a) f. (f (u::'b) v))"
   apply (rule ext)+
-  apply (simp add: list_case_compute_def)
+  apply (simp add: case_list_compute_def)
   done
 
-lemmas compute_list_case = list_case_compute list_case_compute_empty list_case_compute_cons
+lemmas compute_case_list = case_list_compute case_list_compute_empty case_list_compute_cons
 
 (*** compute_list_nth ***)
 (* Of course, you will need computation with nats for this to work \<dots> *)
@@ -122,7 +122,7 @@
   
 (*** compute_list ***)
 
-lemmas compute_list = compute_list_case compute_list_length compute_list_nth
+lemmas compute_list = compute_case_list compute_list_length compute_list_nth
 
 (*** compute_let ***)
 
--- a/src/HOL/Matrix_LP/matrixlp.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Matrix_LP/matrixlp.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -11,7 +11,7 @@
 structure MatrixLP : MATRIX_LP =
 struct
 
-val compute_thms = ComputeHOL.prep_thms @{thms "ComputeHOL.compute_list_case" "ComputeHOL.compute_let"
+val compute_thms = ComputeHOL.prep_thms @{thms "ComputeHOL.compute_case_list" "ComputeHOL.compute_let"
   "ComputeHOL.compute_if" "ComputeFloat.arith" "SparseMatrix.sparse_row_matrix_arith_simps"
   "ComputeHOL.compute_bool" "ComputeHOL.compute_pair"
   "SparseMatrix.sorted_sp_simps"
--- a/src/HOL/Metis_Examples/Clausification.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Metis_Examples/Clausification.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -138,7 +138,7 @@
 by (metis split_list_last_prop[where P = P] in_set_conv_decomp)
 
 lemma ex_tl: "EX ys. tl ys = xs"
-using tl.simps(2) by fast
+using list.sel(3) by fast
 
 lemma "(\<exists>ys\<Colon>nat list. tl ys = xs) \<and> (\<exists>bs\<Colon>int list. tl bs = as)"
 by (metis ex_tl)
--- a/src/HOL/MicroJava/Comp/CorrCompTp.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/MicroJava/Comp/CorrCompTp.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -80,7 +80,7 @@
 apply (rule allI)
 apply (drule_tac x="a # ys" in spec)
 apply (simp only: rev.simps append_assoc append_Cons append_Nil
-  map.simps map_of.simps map_upds_Cons hd.simps tl.simps)
+  map.simps map_of.simps map_upds_Cons list.sel)
 done
 
 lemma map_of_as_map_upds: "map_of (rev xs) = empty ((map fst xs) [\<mapsto>] (map snd xs))"
--- a/src/HOL/MicroJava/Comp/Index.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/MicroJava/Comp/Index.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -54,6 +54,7 @@
 apply (simp only: index_def gjmb_plns_def)
 apply (case_tac "gmb G C S" rule: prod.exhaust)
 apply (simp add: galldefs del: set_append map_append)
+apply (rename_tac a b)
 apply (case_tac b, simp add: gmb_def gjmb_lvs_def del: set_append map_append)
 apply (intro strip)
 apply (simp del: set_append map_append)
@@ -73,6 +74,7 @@
           locvars_xstate G C S (Norm (h, l(x\<mapsto>val)))"
 apply (simp only: locvars_xstate_def locvars_locals_def index_def)
 apply (case_tac "gmb G C S" rule: prod.exhaust, simp)
+apply (rename_tac a b)
 apply (case_tac b, simp)
 apply (rule conjI)
 apply (simp add: gl_def)
--- a/src/HOL/Multivariate_Analysis/Integration.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Multivariate_Analysis/Integration.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -3870,6 +3870,7 @@
     apply rule
     apply (erule_tac x="Some y" in allE)
     defer
+    apply (rename_tac x)
     apply (erule_tac x="Some x" in allE)
     apply auto
     done
--- a/src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -766,7 +766,7 @@
     bs: "set bs = Basis" "distinct bs"
     by (metis finite_distinct_list)
   from nonempty_Basis s obtain j where j: "j \<in> Basis" "s j \<in> S" by blast
-  def y \<equiv> "list_rec
+  def y \<equiv> "rec_list
     (s j)
     (\<lambda>j _ Y. (\<Sum>i\<in>Basis. (if i = j then s i \<bullet> i else Y \<bullet> i) *\<^sub>R i))"
   have "x = (\<Sum>i\<in>Basis. (if i \<in> set bs then s i \<bullet> i else s j \<bullet> i) *\<^sub>R i)"
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -2776,7 +2776,7 @@
       unfolding s_def by (auto intro: someI2_ex)
   }
   note s = this
-  def r \<equiv> "nat_rec (s 0 0) s"
+  def r \<equiv> "rec_nat (s 0 0) s"
   have "subseq r"
     by (auto simp: r_def s subseq_Suc_iff)
   moreover
@@ -3376,7 +3376,7 @@
       unfolding s_def by (auto intro: someI2_ex)
   }
   note s = this
-  def r \<equiv> "nat_rec (s 0 0) s"
+  def r \<equiv> "rec_nat (s 0 0) s"
   have "subseq r"
     by (auto simp: r_def s subseq_Suc_iff)
   moreover
@@ -3953,7 +3953,7 @@
       }
       note B = this
 
-      def F \<equiv> "nat_rec (B 0 UNIV) B"
+      def F \<equiv> "rec_nat (B 0 UNIV) B"
       {
         fix n
         have "infinite {i. f i \<in> F n}"
@@ -3974,7 +3974,7 @@
           by (simp add: set_eq_iff not_le conj_commute)
       qed
 
-      def t \<equiv> "nat_rec (sel 0 0) (\<lambda>n i. sel (Suc n) i)"
+      def t \<equiv> "rec_nat (sel 0 0) (\<lambda>n i. sel (Suc n) i)"
       have "subseq t"
         unfolding subseq_Suc_iff by (simp add: t_def sel)
       moreover have "\<forall>i. (f \<circ> t) i \<in> s"
--- a/src/HOL/Nat.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Nat.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -71,33 +71,81 @@
 lemma Suc_Rep_inject': "Suc_Rep x = Suc_Rep y \<longleftrightarrow> x = y"
   by (rule iffI, rule Suc_Rep_inject) simp_all
 
+lemma nat_induct0:
+  fixes n
+  assumes "P 0" and "\<And>n. P n \<Longrightarrow> P (Suc n)"
+  shows "P n"
+using assms
+apply (unfold Zero_nat_def Suc_def)
+apply (rule Rep_Nat_inverse [THEN subst]) -- {* types force good instantiation *}
+apply (erule Nat_Rep_Nat [THEN Nat.induct])
+apply (iprover elim: Nat_Abs_Nat_inverse [THEN subst])
+done
+
+wrap_free_constructors ["0 \<Colon> nat", Suc] case_nat [=] [[], [pred]]
+  apply atomize_elim
+  apply (rename_tac n, induct_tac n rule: nat_induct0, auto)
+ apply (simp add: Suc_def Nat_Abs_Nat_inject Nat_Rep_Nat Suc_RepI
+   Suc_Rep_inject' Rep_Nat_inject)
+apply (simp only: Suc_not_Zero)
+done
+
+-- {* Avoid name clashes by prefixing the output of @{text rep_datatype} with @{text old}. *}
+setup {* Sign.mandatory_path "old" *}
+
 rep_datatype "0 \<Colon> nat" Suc
-  apply (unfold Zero_nat_def Suc_def)
-  apply (rule Rep_Nat_inverse [THEN subst]) -- {* types force good instantiation *}
-   apply (erule Nat_Rep_Nat [THEN Nat.induct])
-   apply (iprover elim: Nat_Abs_Nat_inverse [THEN subst])
-    apply (simp_all add: Nat_Abs_Nat_inject Nat_Rep_Nat
-      Suc_RepI Zero_RepI Suc_Rep_not_Zero_Rep
-      Suc_Rep_not_Zero_Rep [symmetric]
-      Suc_Rep_inject' Rep_Nat_inject)
-  done
+  apply (erule nat_induct0, assumption)
+ apply (rule nat.inject)
+apply (rule nat.distinct(1))
+done
+
+setup {* Sign.parent_path *}
+
+-- {* But erase the prefix for properties that are not generated by @{text wrap_free_constructors}. *}
+setup {* Sign.mandatory_path "nat" *}
+
+declare
+  old.nat.inject[iff del]
+  old.nat.distinct(1)[simp del, induct_simp del]
+
+lemmas induct = old.nat.induct
+lemmas inducts = old.nat.inducts
+lemmas recs = old.nat.recs
+lemmas cases = nat.case
+lemmas simps = nat.inject nat.distinct nat.case old.nat.recs
+
+setup {* Sign.parent_path *}
+
+abbreviation rec_nat :: "'a \<Rightarrow> (nat \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a" where
+  "rec_nat \<equiv> old.rec_nat"
+
+declare nat.sel[code del]
+
+hide_const Nat.pred -- {* hide everything related to the selector *}
+hide_fact
+  nat.case_eq_if
+  nat.collapse
+  nat.expand
+  nat.sel
+  nat.sel_exhaust
+  nat.sel_split
+  nat.sel_split_asm
+
+lemma nat_exhaust [case_names 0 Suc, cases type: nat]:
+  -- {* for backward compatibility -- names of variables differ *}
+  "(y = 0 \<Longrightarrow> P) \<Longrightarrow> (\<And>nat. y = Suc nat \<Longrightarrow> P) \<Longrightarrow> P"
+by (rule old.nat.exhaust)
 
 lemma nat_induct [case_names 0 Suc, induct type: nat]:
   -- {* for backward compatibility -- names of variables differ *}
   fixes n
-  assumes "P 0"
-    and "\<And>n. P n \<Longrightarrow> P (Suc n)"
+  assumes "P 0" and "\<And>n. P n \<Longrightarrow> P (Suc n)"
   shows "P n"
-  using assms by (rule nat.induct)
-
-declare nat.exhaust [case_names 0 Suc, cases type: nat]
+using assms by (rule nat.induct)
 
-lemmas nat_rec_0 = nat.recs(1)
-  and nat_rec_Suc = nat.recs(2)
-
-lemmas nat_case_0 = nat.cases(1)
-  and nat_case_Suc = nat.cases(2)
-   
+hide_fact
+  nat_exhaust
+  nat_induct0
 
 text {* Injectiveness and distinctness lemmas *}
 
@@ -632,14 +680,6 @@
 
 lemmas not_less_simps = not_less_less_Suc_eq le_less_Suc_eq
 
-text {* These two rules ease the use of primitive recursion.
-NOTE USE OF @{text "=="} *}
-lemma def_nat_rec_0: "(!!n. f n == nat_rec c h n) ==> f 0 = c"
-by simp
-
-lemma def_nat_rec_Suc: "(!!n. f n == nat_rec c h n) ==> f (Suc n) = h n (f n)"
-by simp
-
 lemma not0_implies_Suc: "n \<noteq> 0 ==> \<exists>m. n = Suc m"
 by (cases n) simp_all
 
@@ -1905,13 +1945,13 @@
 qed
 
 
-subsection {* aliasses *}
+subsection {* aliases *}
 
 lemma nat_mult_1: "(1::nat) * n = n"
-  by simp
+  by (rule mult_1_left)
  
 lemma nat_mult_1_right: "n * (1::nat) = n"
-  by simp
+  by (rule mult_1_right)
 
 
 subsection {* size of a datatype value *}
@@ -1928,4 +1968,3 @@
 hide_const (open) of_nat_aux
 
 end
-
--- a/src/HOL/Nitpick.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Nitpick.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -96,16 +96,16 @@
 apply (erule contrapos_np)
 by (rule someI)
 
-lemma unit_case_unfold [nitpick_unfold]:
-"unit_case x u \<equiv> x"
+lemma case_unit_unfold [nitpick_unfold]:
+"case_unit x u \<equiv> x"
 apply (subgoal_tac "u = ()")
  apply (simp only: unit.cases)
 by simp
 
 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 unit_case_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
--- a/src/HOL/Nitpick_Examples/Core_Nits.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Nitpick_Examples/Core_Nits.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -925,18 +925,6 @@
 nitpick [card = 2, expect = none]
 by auto
 
-lemma "bool_rec x y True = x"
-nitpick [card = 2, expect = none]
-by auto
-
-lemma "bool_rec x y False = y"
-nitpick [card = 2, expect = none]
-by auto
-
-lemma "(x\<Colon>bool) = bool_rec x x True"
-nitpick [card = 2, expect = none]
-by auto
-
 lemma "x = (case (x, y) of (x', y') \<Rightarrow> x')"
 nitpick [expect = none]
 sorry
--- a/src/HOL/Nitpick_Examples/Refute_Nits.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Nitpick_Examples/Refute_Nits.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -519,18 +519,6 @@
 nitpick [expect = genuine]
 oops
 
-lemma "prod_rec f p = f (fst p) (snd p)"
-nitpick [expect = none]
-by (case_tac p) auto
-
-lemma "prod_rec f (a, b) = f a b"
-nitpick [expect = none]
-by auto
-
-lemma "P (prod_rec f x)"
-nitpick [expect = genuine]
-oops
-
 lemma "P (case x of Pair a b \<Rightarrow> f a b)"
 nitpick [expect = genuine]
 oops
@@ -575,15 +563,6 @@
 nitpick [expect = genuine]
 oops
 
-lemma "unit_rec u x = u"
-nitpick [expect = none]
-apply simp
-done
-
-lemma "P (unit_rec u x)"
-nitpick [expect = genuine]
-oops
-
 lemma "P (case x of () \<Rightarrow> u)"
 nitpick [expect = genuine]
 oops
@@ -606,17 +585,17 @@
 nitpick [expect = genuine]
 oops
 
-lemma "option_rec n s None = n"
+lemma "rec_option n s None = n"
 nitpick [expect = none]
 apply simp
 done
 
-lemma "option_rec n s (Some x) = s x"
+lemma "rec_option n s (Some x) = s x"
 nitpick [expect = none]
 apply simp
 done
 
-lemma "P (option_rec n s x)"
+lemma "P (rec_option n s x)"
 nitpick [expect = genuine]
 oops
 
@@ -646,20 +625,6 @@
 nitpick [expect = genuine]
 oops
 
-lemma "sum_rec l r (Inl x) = l x"
-nitpick [expect = none]
-apply simp
-done
-
-lemma "sum_rec l r (Inr x) = r x"
-nitpick [expect = none]
-apply simp
-done
-
-lemma "P (sum_rec l r x)"
-nitpick [expect = genuine]
-oops
-
 lemma "P (case x of Inl a \<Rightarrow> l a | Inr b \<Rightarrow> r b)"
 nitpick [expect = genuine]
 oops
@@ -684,17 +649,17 @@
 nitpick [expect = genuine]
 oops
 
-lemma "T1_rec a b A = a"
+lemma "rec_T1 a b A = a"
 nitpick [expect = none]
 apply simp
 done
 
-lemma "T1_rec a b B = b"
+lemma "rec_T1 a b B = b"
 nitpick [expect = none]
 apply simp
 done
 
-lemma "P (T1_rec a b x)"
+lemma "P (rec_T1 a b x)"
 nitpick [expect = genuine]
 oops
 
@@ -716,17 +681,17 @@
 nitpick [expect = genuine]
 oops
 
-lemma "T2_rec c d (C x) = c x"
+lemma "rec_T2 c d (C x) = c x"
 nitpick [expect = none]
 apply simp
 done
 
-lemma "T2_rec c d (D x) = d x"
+lemma "rec_T2 c d (D x) = d x"
 nitpick [expect = none]
 apply simp
 done
 
-lemma "P (T2_rec c d x)"
+lemma "P (rec_T2 c d x)"
 nitpick [expect = genuine]
 oops
 
@@ -748,12 +713,12 @@
 nitpick [expect = genuine]
 oops
 
-lemma "T3_rec e (E x) = e x"
+lemma "rec_T3 e (E x) = e x"
 nitpick [card = 1\<emdash>4, expect = none]
 apply simp
 done
 
-lemma "P (T3_rec e x)"
+lemma "P (rec_T3 e x)"
 nitpick [expect = genuine]
 oops
 
@@ -781,17 +746,17 @@
 nitpick [card = 1\<emdash>7, expect = none]
 oops
 
-lemma "nat_rec zero suc 0 = zero"
+lemma "rec_nat zero suc 0 = zero"
 nitpick [expect = none]
 apply simp
 done
 
-lemma "nat_rec zero suc (Suc x) = suc x (nat_rec zero suc x)"
+lemma "rec_nat zero suc (Suc x) = suc x (rec_nat zero suc x)"
 nitpick [expect = none]
 apply simp
 done
 
-lemma "P (nat_rec zero suc x)"
+lemma "P (rec_nat zero suc x)"
 nitpick [expect = genuine]
 oops
 
@@ -813,17 +778,17 @@
 nitpick [expect = genuine]
 oops
 
-lemma "list_rec nil cons [] = nil"
+lemma "rec_list nil cons [] = nil"
 nitpick [card = 1\<emdash>5, expect = none]
 apply simp
 done
 
-lemma "list_rec nil cons (x#xs) = cons x xs (list_rec nil cons xs)"
+lemma "rec_list nil cons (x#xs) = cons x xs (rec_list nil cons xs)"
 nitpick [card = 1\<emdash>5, expect = none]
 apply simp
 done
 
-lemma "P (list_rec nil cons xs)"
+lemma "P (rec_list nil cons xs)"
 nitpick [expect = genuine]
 oops
 
@@ -853,22 +818,22 @@
 nitpick [expect = genuine]
 oops
 
-lemma "BitList_rec nil bit0 bit1 BitListNil = nil"
+lemma "rec_BitList nil bit0 bit1 BitListNil = nil"
 nitpick [expect = none]
 apply simp
 done
 
-lemma "BitList_rec nil bit0 bit1 (Bit0 xs) = bit0 xs (BitList_rec nil bit0 bit1 xs)"
+lemma "rec_BitList nil bit0 bit1 (Bit0 xs) = bit0 xs (rec_BitList nil bit0 bit1 xs)"
 nitpick [expect = none]
 apply simp
 done
 
-lemma "BitList_rec nil bit0 bit1 (Bit1 xs) = bit1 xs (BitList_rec nil bit0 bit1 xs)"
+lemma "rec_BitList nil bit0 bit1 (Bit1 xs) = bit1 xs (rec_BitList nil bit0 bit1 xs)"
 nitpick [expect = none]
 apply simp
 done
 
-lemma "P (BitList_rec nil bit0 bit1 x)"
+lemma "P (rec_BitList nil bit0 bit1 x)"
 nitpick [expect = genuine]
 oops
 
@@ -886,17 +851,17 @@
 nitpick [expect = genuine]
 oops
 
-lemma "BinTree_rec l n (Leaf x) = l x"
+lemma "rec_BinTree l n (Leaf x) = l x"
 nitpick [expect = none]
 apply simp
 done
 
-lemma "BinTree_rec l n (Node x y) = n x y (BinTree_rec l n x) (BinTree_rec l n y)"
+lemma "rec_BinTree l n (Node x y) = n x y (rec_BinTree l n x) (rec_BinTree l n y)"
 nitpick [card = 1\<emdash>5, expect = none]
 apply simp
 done
 
-lemma "P (BinTree_rec l n x)"
+lemma "P (rec_BinTree l n x)"
 nitpick [expect = genuine]
 oops
 
@@ -929,17 +894,17 @@
 nitpick [expect = genuine]
 oops
 
-lemma "aexp_bexp_rec_1 number ite equal (Number x) = number x"
+lemma "rec_aexp_bexp_1 number ite equal (Number x) = number x"
 nitpick [card = 1\<emdash>3, expect = none]
 apply simp
 done
 
-lemma "aexp_bexp_rec_1 number ite equal (ITE x y z) = ite x y z (aexp_bexp_rec_2 number ite equal x) (aexp_bexp_rec_1 number ite equal y) (aexp_bexp_rec_1 number ite equal z)"
+lemma "rec_aexp_bexp_1 number ite equal (ITE x y z) = ite x y z (rec_aexp_bexp_2 number ite equal x) (rec_aexp_bexp_1 number ite equal y) (rec_aexp_bexp_1 number ite equal z)"
 nitpick [card = 1\<emdash>3, expect = none]
 apply simp
 done
 
-lemma "P (aexp_bexp_rec_1 number ite equal x)"
+lemma "P (rec_aexp_bexp_1 number ite equal x)"
 nitpick [expect = genuine]
 oops
 
@@ -947,12 +912,12 @@
 nitpick [expect = genuine]
 oops
 
-lemma "aexp_bexp_rec_2 number ite equal (Equal x y) = equal x y (aexp_bexp_rec_1 number ite equal x) (aexp_bexp_rec_1 number ite equal y)"
+lemma "rec_aexp_bexp_2 number ite equal (Equal x y) = equal x y (rec_aexp_bexp_1 number ite equal x) (rec_aexp_bexp_1 number ite equal y)"
 nitpick [card = 1\<emdash>3, expect = none]
 apply simp
 done
 
-lemma "P (aexp_bexp_rec_2 number ite equal x)"
+lemma "P (rec_aexp_bexp_2 number ite equal x)"
 nitpick [expect = genuine]
 oops
 
@@ -1007,41 +972,41 @@
 nitpick [expect = genuine]
 oops
 
-lemma "X_Y_rec_1 a b c d e f A = a"
+lemma "rec_X_Y_1 a b c d e f A = a"
 nitpick [card = 1\<emdash>5, expect = none]
 apply simp
 done
 
-lemma "X_Y_rec_1 a b c d e f (B x) = b x (X_Y_rec_1 a b c d e f x)"
+lemma "rec_X_Y_1 a b c d e f (B x) = b x (rec_X_Y_1 a b c d e f x)"
 nitpick [card = 1\<emdash>5, expect = none]
 apply simp
 done
 
-lemma "X_Y_rec_1 a b c d e f (C y) = c y (X_Y_rec_2 a b c d e f y)"
+lemma "rec_X_Y_1 a b c d e f (C y) = c y (rec_X_Y_2 a b c d e f y)"
 nitpick [card = 1\<emdash>5, expect = none]
 apply simp
 done
 
-lemma "X_Y_rec_2 a b c d e f (D x) = d x (X_Y_rec_1 a b c d e f x)"
+lemma "rec_X_Y_2 a b c d e f (D x) = d x (rec_X_Y_1 a b c d e f x)"
 nitpick [card = 1\<emdash>5, expect = none]
 apply simp
 done
 
-lemma "X_Y_rec_2 a b c d e f (E y) = e y (X_Y_rec_2 a b c d e f y)"
+lemma "rec_X_Y_2 a b c d e f (E y) = e y (rec_X_Y_2 a b c d e f y)"
 nitpick [card = 1\<emdash>5, expect = none]
 apply simp
 done
 
-lemma "X_Y_rec_2 a b c d e f F = f"
+lemma "rec_X_Y_2 a b c d e f F = f"
 nitpick [card = 1\<emdash>5, expect = none]
 apply simp
 done
 
-lemma "P (X_Y_rec_1 a b c d e f x)"
+lemma "P (rec_X_Y_1 a b c d e f x)"
 nitpick [expect = genuine]
 oops
 
-lemma "P (X_Y_rec_2 a b c d e f y)"
+lemma "P (rec_X_Y_2 a b c d e f y)"
 nitpick [expect = genuine]
 oops
 
@@ -1063,45 +1028,45 @@
 nitpick [expect = genuine]
 oops
 
-lemma "XOpt_rec_1 cx dx n1 s1 n2 s2 (CX x) = cx x (XOpt_rec_2 cx dx n1 s1 n2 s2 x)"
+lemma "rec_XOpt_1 cx dx n1 s1 n2 s2 (CX x) = cx x (rec_XOpt_2 cx dx n1 s1 n2 s2 x)"
 nitpick [card = 1\<emdash>5, expect = none]
 apply simp
 done
 
-lemma "XOpt_rec_1 cx dx n1 s1 n2 s2 (DX x) = dx x (\<lambda>b. XOpt_rec_3 cx dx n1 s1 n2 s2 (x b))"
+lemma "rec_XOpt_1 cx dx n1 s1 n2 s2 (DX x) = dx x (\<lambda>b. rec_XOpt_3 cx dx n1 s1 n2 s2 (x b))"
 nitpick [card = 1\<emdash>3, expect = none]
 apply simp
 done
 
-lemma "XOpt_rec_2 cx dx n1 s1 n2 s2 None = n1"
+lemma "rec_XOpt_2 cx dx n1 s1 n2 s2 None = n1"
 nitpick [card = 1\<emdash>4, expect = none]
 apply simp
 done
 
-lemma "XOpt_rec_2 cx dx n1 s1 n2 s2 (Some x) = s1 x (XOpt_rec_1 cx dx n1 s1 n2 s2 x)"
+lemma "rec_XOpt_2 cx dx n1 s1 n2 s2 (Some x) = s1 x (rec_XOpt_1 cx dx n1 s1 n2 s2 x)"
 nitpick [card = 1\<emdash>4, expect = none]
 apply simp
 done
 
-lemma "XOpt_rec_3 cx dx n1 s1 n2 s2 None = n2"
+lemma "rec_XOpt_3 cx dx n1 s1 n2 s2 None = n2"
 nitpick [card = 1\<emdash>4, expect = none]
 apply simp
 done
 
-lemma "XOpt_rec_3 cx dx n1 s1 n2 s2 (Some x) = s2 x (XOpt_rec_1 cx dx n1 s1 n2 s2 x)"
+lemma "rec_XOpt_3 cx dx n1 s1 n2 s2 (Some x) = s2 x (rec_XOpt_1 cx dx n1 s1 n2 s2 x)"
 nitpick [card = 1\<emdash>4, expect = none]
 apply simp
 done
 
-lemma "P (XOpt_rec_1 cx dx n1 s1 n2 s2 x)"
+lemma "P (rec_XOpt_1 cx dx n1 s1 n2 s2 x)"
 nitpick [expect = genuine]
 oops
 
-lemma "P (XOpt_rec_2 cx dx n1 s1 n2 s2 x)"
+lemma "P (rec_XOpt_2 cx dx n1 s1 n2 s2 x)"
 nitpick [expect = genuine]
 oops
 
-lemma "P (XOpt_rec_3 cx dx n1 s1 n2 s2 x)"
+lemma "P (rec_XOpt_3 cx dx n1 s1 n2 s2 x)"
 nitpick [expect = genuine]
 oops
 
@@ -1119,26 +1084,26 @@
 nitpick [expect = genuine]
 oops
 
-lemma "YOpt_rec_1 cy n s (CY x) = cy x (YOpt_rec_2 cy n s x)"
+lemma "rec_YOpt_1 cy n s (CY x) = cy x (rec_YOpt_2 cy n s x)"
 nitpick [card = 1\<emdash>2, expect = none]
 apply simp
 done
 
-lemma "YOpt_rec_2 cy n s None = n"
+lemma "rec_YOpt_2 cy n s None = n"
 nitpick [card = 1\<emdash>2, expect = none]
 apply simp
 done
 
-lemma "YOpt_rec_2 cy n s (Some x) = s x (\<lambda>a. YOpt_rec_1 cy n s (x a))"
+lemma "rec_YOpt_2 cy n s (Some x) = s x (\<lambda>a. rec_YOpt_1 cy n s (x a))"
 nitpick [card = 1\<emdash>2, expect = none]
 apply simp
 done
 
-lemma "P (YOpt_rec_1 cy n s x)"
+lemma "P (rec_YOpt_1 cy n s x)"
 nitpick [expect = genuine]
 oops
 
-lemma "P (YOpt_rec_2 cy n s x)"
+lemma "P (rec_YOpt_2 cy n s x)"
 nitpick [expect = genuine]
 oops
 
@@ -1156,26 +1121,26 @@
 nitpick [expect = genuine]
 oops
 
-lemma "Trie_rec_1 tr nil cons (TR x) = tr x (Trie_rec_2 tr nil cons x)"
+lemma "rec_Trie_1 tr nil cons (TR x) = tr x (rec_Trie_2 tr nil cons x)"
 nitpick [card = 1\<emdash>4, expect = none]
 apply simp
 done
 
-lemma "Trie_rec_2 tr nil cons [] = nil"
+lemma "rec_Trie_2 tr nil cons [] = nil"
 nitpick [card = 1\<emdash>4, expect = none]
 apply simp
 done
 
-lemma "Trie_rec_2 tr nil cons (x#xs) = cons x xs (Trie_rec_1 tr nil cons x) (Trie_rec_2 tr nil cons xs)"
+lemma "rec_Trie_2 tr nil cons (x#xs) = cons x xs (rec_Trie_1 tr nil cons x) (rec_Trie_2 tr nil cons xs)"
 nitpick [card = 1\<emdash>4, expect = none]
 apply simp
 done
 
-lemma "P (Trie_rec_1 tr nil cons x)"
+lemma "P (rec_Trie_1 tr nil cons x)"
 nitpick [card = 1, expect = genuine]
 oops
 
-lemma "P (Trie_rec_2 tr nil cons x)"
+lemma "P (rec_Trie_2 tr nil cons x)"
 nitpick [card = 1, expect = genuine]
 oops
 
@@ -1193,17 +1158,17 @@
 nitpick [expect = genuine]
 oops
 
-lemma "InfTree_rec leaf node Leaf = leaf"
+lemma "rec_InfTree leaf node Leaf = leaf"
 nitpick [card = 1\<emdash>3, expect = none]
 apply simp
 done
 
-lemma "InfTree_rec leaf node (Node x) = node x (\<lambda>n. InfTree_rec leaf node (x n))"
+lemma "rec_InfTree leaf node (Node x) = node x (\<lambda>n. rec_InfTree leaf node (x n))"
 nitpick [card = 1\<emdash>3, expect = none]
 apply simp
 done
 
-lemma "P (InfTree_rec leaf node x)"
+lemma "P (rec_InfTree leaf node x)"
 nitpick [expect = genuine]
 oops
 
@@ -1222,22 +1187,22 @@
 nitpick [card 'a = 4, card "'a lambda" = 5, expect = genuine]
 oops
 
-lemma "lambda_rec var app lam (Var x) = var x"
+lemma "rec_lambda var app lam (Var x) = var x"
 nitpick [card = 1\<emdash>3, expect = none]
 apply simp
 done
 
-lemma "lambda_rec var app lam (App x y) = app x y (lambda_rec var app lam x) (lambda_rec var app lam y)"
+lemma "rec_lambda var app lam (App x y) = app x y (rec_lambda var app lam x) (rec_lambda var app lam y)"
 nitpick [card = 1\<emdash>3, expect = none]
 apply simp
 done
 
-lemma "lambda_rec var app lam (Lam x) = lam x (\<lambda>a. lambda_rec var app lam (x a))"
+lemma "rec_lambda var app lam (Lam x) = lam x (\<lambda>a. rec_lambda var app lam (x a))"
 nitpick [card = 1\<emdash>3, expect = none]
 apply simp
 done
 
-lemma "P (lambda_rec v a l x)"
+lemma "P (rec_lambda v a l x)"
 nitpick [expect = genuine]
 oops
 
@@ -1258,40 +1223,40 @@
 nitpick [expect = genuine]
 oops
 
-lemma "U_rec_1 e c d nil cons (E x) = e x (U_rec_2 e c d nil cons x)"
+lemma "rec_U_1 e c d nil cons (E x) = e x (rec_U_2 e c d nil cons x)"
 nitpick [card = 1\<emdash>3, expect = none]
 apply simp
 done
 
-lemma "U_rec_2 e c d nil cons (C x) = c x"
+lemma "rec_U_2 e c d nil cons (C x) = c x"
 nitpick [card = 1\<emdash>3, expect = none]
 apply simp
 done
 
-lemma "U_rec_2 e c d nil cons (D x) = d x (U_rec_3 e c d nil cons x)"
+lemma "rec_U_2 e c d nil cons (D x) = d x (rec_U_3 e c d nil cons x)"
 nitpick [card = 1\<emdash>3, expect = none]
 apply simp
 done
 
-lemma "U_rec_3 e c d nil cons [] = nil"
+lemma "rec_U_3 e c d nil cons [] = nil"
 nitpick [card = 1\<emdash>3, expect = none]
 apply simp
 done
 
-lemma "U_rec_3 e c d nil cons (x#xs) = cons x xs (U_rec_1 e c d nil cons x) (U_rec_3 e c d nil cons xs)"
+lemma "rec_U_3 e c d nil cons (x#xs) = cons x xs (rec_U_1 e c d nil cons x) (rec_U_3 e c d nil cons xs)"
 nitpick [card = 1\<emdash>3, expect = none]
 apply simp
 done
 
-lemma "P (U_rec_1 e c d nil cons x)"
+lemma "P (rec_U_1 e c d nil cons x)"
 nitpick [expect = genuine]
 oops
 
-lemma "P (U_rec_2 e c d nil cons x)"
+lemma "P (rec_U_2 e c d nil cons x)"
 nitpick [card = 1, expect = genuine]
 oops
 
-lemma "P (U_rec_3 e c d nil cons x)"
+lemma "P (rec_U_3 e c d nil cons x)"
 nitpick [card = 1, expect = genuine]
 oops
 
--- a/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -164,7 +164,7 @@
 by (rule Rep_Nat_inverse)
 
 lemma "Abs_list (Rep_list a) = a"
-nitpick [card = 1\<emdash>2, expect = none]
+(* nitpick [card = 1\<emdash>2, expect = none] FIXME *)
 by (rule Rep_list_inverse)
 
 record point =
--- a/src/HOL/Nominal/Examples/Class2.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Nominal/Examples/Class2.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -2881,7 +2881,7 @@
 done
 
 termination
-apply(relation "measure (sum_case (size\<circ>fst) (size\<circ>fst))")
+apply(relation "measure (case_sum (size\<circ>fst) (size\<circ>fst))")
 apply(simp_all)
 done
 
--- a/src/HOL/Nominal/Examples/Class3.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Nominal/Examples/Class3.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -2311,6 +2311,7 @@
 apply(rule_tac my_wf_induct_triple[OF a])
 apply(case_tac x rule: prod.exhaust)
 apply(simp)
+apply(rename_tac p a b)
 apply(case_tac b)
 apply(simp)
 apply(rule b)
@@ -3618,7 +3619,7 @@
 apply(simp add: fresh_atm)
 done
 
-lemma option_case_eqvt1[eqvt_force]:
+lemma case_option_eqvt1[eqvt_force]:
   fixes pi1::"name prm"
   and   pi2::"coname prm"
   and   B::"(name\<times>trm) option"
@@ -3635,7 +3636,7 @@
 apply(perm_simp)
 done
 
-lemma option_case_eqvt2[eqvt_force]:
+lemma case_option_eqvt2[eqvt_force]:
   fixes pi1::"name prm"
   and   pi2::"coname prm"
   and   B::"(coname\<times>trm) option"
--- a/src/HOL/Nominal/Examples/Fsub.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Nominal/Examples/Fsub.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -167,7 +167,8 @@
   assumes a: "X\<in>(ty_dom \<Gamma>)" 
   shows "\<exists>T.(TVarB X T)\<in>set \<Gamma>"
   using a 
-  apply (induct \<Gamma>, auto) 
+  apply (induct \<Gamma>, auto)
+  apply (rename_tac a \<Gamma>') 
   apply (subgoal_tac "\<exists>T.(TVarB X T=a)")
   apply (auto)
   apply (auto simp add: ty_binding_existence)
--- a/src/HOL/Nominal/Examples/W.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Nominal/Examples/W.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -1,5 +1,5 @@
 theory W
-imports Nominal
+imports "../Nominal"
 begin
 
 text {* Example for strong induction rules avoiding sets of atoms. *}
@@ -388,6 +388,7 @@
   shows "supp \<Gamma> = set (ftv \<Gamma>)"
 apply (induct \<Gamma>)
 apply (simp_all add: supp_list_nil supp_list_cons)
+apply (rename_tac a \<Gamma>')
 apply (case_tac a)
 apply (simp add: supp_prod supp_atm ftv_tyS)
 done
@@ -443,6 +444,7 @@
 using asm
 apply(induct Xs)
 apply(simp)
+apply(rename_tac a Xs')
 apply(case_tac "X=a")
 apply(simp add: abs_fresh)
 apply(simp add: abs_fresh)
--- a/src/HOL/Num.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Num.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -1050,24 +1050,24 @@
   "min (numeral k) (Suc n) = Suc (min (pred_numeral k) n)"
   by (simp add: numeral_eq_Suc)
 
-text {* For @{term nat_case} and @{term nat_rec}. *}
+text {* For @{term case_nat} and @{term rec_nat}. *}
 
-lemma nat_case_numeral [simp]:
-  "nat_case a f (numeral v) = (let pv = pred_numeral v in f pv)"
+lemma case_nat_numeral [simp]:
+  "case_nat a f (numeral v) = (let pv = pred_numeral v in f pv)"
   by (simp add: numeral_eq_Suc)
 
-lemma nat_case_add_eq_if [simp]:
-  "nat_case a f ((numeral v) + n) = (let pv = pred_numeral v in f (pv + n))"
+lemma case_nat_add_eq_if [simp]:
+  "case_nat a f ((numeral v) + n) = (let pv = pred_numeral v in f (pv + n))"
   by (simp add: numeral_eq_Suc)
 
-lemma nat_rec_numeral [simp]:
-  "nat_rec a f (numeral v) =
-    (let pv = pred_numeral v in f pv (nat_rec a f pv))"
+lemma rec_nat_numeral [simp]:
+  "rec_nat a f (numeral v) =
+    (let pv = pred_numeral v in f pv (rec_nat a f pv))"
   by (simp add: numeral_eq_Suc Let_def)
 
-lemma nat_rec_add_eq_if [simp]:
-  "nat_rec a f (numeral v + n) =
-    (let pv = pred_numeral v in f (pv + n) (nat_rec a f (pv + n)))"
+lemma rec_nat_add_eq_if [simp]:
+  "rec_nat a f (numeral v + n) =
+    (let pv = pred_numeral v in f (pv + n) (rec_nat a f (pv + n)))"
   by (simp add: numeral_eq_Suc Let_def)
 
 text {* Case analysis on @{term "n < 2"} *}
--- a/src/HOL/Option.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Option.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -5,10 +5,33 @@
 header {* Datatype option *}
 
 theory Option
-imports Datatype Finite_Set
+imports BNF_LFP Datatype Finite_Set
 begin
 
-datatype 'a option = None | Some 'a
+datatype_new 'a option =
+    =: None
+  | Some (the: 'a)
+
+datatype_new_compat option
+
+lemma [case_names None Some, cases type: option]:
+  -- {* for backward compatibility -- names of variables differ *}
+  "(y = None \<Longrightarrow> P) \<Longrightarrow> (\<And>a. y = Some a \<Longrightarrow> P) \<Longrightarrow> P"
+by (rule option.exhaust)
+
+lemma [case_names None Some, induct type: option]:
+  -- {* for backward compatibility -- names of variables differ *}
+  "P None \<Longrightarrow> (\<And>option. P (Some option)) \<Longrightarrow> P option"
+by (rule option.induct)
+
+-- {* Compatibility *}
+setup {* Sign.mandatory_path "option" *}
+
+lemmas inducts = option.induct
+lemmas recs = option.rec
+lemmas cases = option.case
+
+setup {* Sign.parent_path *}
 
 lemma not_None_eq [iff]: "(x ~= None) = (EX y. x = Some y)"
   by (induct x) auto
@@ -23,7 +46,7 @@
 lemma inj_Some [simp]: "inj_on Some A"
 by (rule inj_onI) simp
 
-lemma option_caseE:
+lemma case_optionE:
   assumes c: "(case x of None => P | Some y => Q y)"
   obtains
     (None) "x = None" and P
@@ -41,9 +64,6 @@
 
 subsubsection {* Operations *}
 
-primrec the :: "'a option => 'a" where
-"the (Some x) = x"
-
 primrec set :: "'a option => 'a set" where
 "set None = {}" |
 "set (Some x) = {x}"
@@ -80,8 +100,8 @@
     "map f (map g opt) = map (f o g) opt"
   by (simp add: map_def split add: option.split)
 
-lemma option_map_o_sum_case [simp]:
-    "map f o sum_case g h = sum_case (map f o g) (map f o h)"
+lemma option_map_o_case_sum [simp]:
+    "map f o case_sum g h = case_sum (map f o g) (map f o h)"
   by (rule ext) (simp split: sum.split)
 
 lemma map_cong: "x = y \<Longrightarrow> (\<And>a. y = Some a \<Longrightarrow> f a = g a) \<Longrightarrow> map f x = map g y"
@@ -104,8 +124,8 @@
   qed
 qed
 
-lemma option_case_map [simp]:
-  "option_case g h (Option.map f x) = option_case g (h \<circ> f) x"
+lemma case_option_map [simp]:
+  "case_option g h (Option.map f x) = case_option g (h \<circ> f) x"
   by (cases x) simp_all
 
 primrec bind :: "'a option \<Rightarrow> ('a \<Rightarrow> 'b option) \<Rightarrow> 'b option" where
--- a/src/HOL/Predicate.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Predicate.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -531,11 +531,11 @@
   by (fact equal_refl)
 
 lemma [code]:
-  "pred_case f P = f (eval P)"
+  "case_pred f P = f (eval P)"
   by (cases P) simp
 
 lemma [code]:
-  "pred_rec f P = f (eval P)"
+  "rec_pred f P = f (eval P)"
   by (cases P) simp
 
 inductive eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where "eq x x"
@@ -722,4 +722,3 @@
 hide_fact (open) null_def member_def
 
 end
-
--- a/src/HOL/Probability/Finite_Product_Measure.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Probability/Finite_Product_Measure.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -442,15 +442,15 @@
   "i \<in> I \<Longrightarrow> f \<in> measurable (M i) N \<Longrightarrow> (\<lambda>x. f (x i)) \<in> measurable (PiM I M) N"
   by simp
 
-lemma measurable_nat_case[measurable (raw)]:
+lemma measurable_case_nat[measurable (raw)]:
   assumes [measurable (raw)]: "i = 0 \<Longrightarrow> f \<in> measurable M N"
     "\<And>j. i = Suc j \<Longrightarrow> (\<lambda>x. g x j) \<in> measurable M N"
-  shows "(\<lambda>x. nat_case (f x) (g x) i) \<in> measurable M N"
+  shows "(\<lambda>x. case_nat (f x) (g x) i) \<in> measurable M N"
   by (cases i) simp_all
 
-lemma measurable_nat_case'[measurable (raw)]:
+lemma measurable_case_nat'[measurable (raw)]:
   assumes fg[measurable]: "f \<in> measurable N M" "g \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
-  shows "(\<lambda>x. nat_case (f x) (g x)) \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
+  shows "(\<lambda>x. case_nat (f x) (g x)) \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
   using fg[THEN measurable_space]
   by (auto intro!: measurable_PiM_single' simp add: space_PiM PiE_iff split: nat.split)
 
@@ -465,7 +465,7 @@
   also have "\<dots> \<in> sets ?P"
     using A j
     by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton])
-  finally show "{\<omega> \<in> space ?P. prod_case (\<lambda>f. fun_upd f i) \<omega> j \<in> A} \<in> sets ?P" .
+  finally show "{\<omega> \<in> space ?P. case_prod (\<lambda>f. fun_upd f i) \<omega> j \<in> A} \<in> sets ?P" .
 qed (auto simp: space_pair_measure space_PiM PiE_def)
 
 lemma measurable_component_update:
@@ -1132,27 +1132,27 @@
 lemma pair_measure_eq_distr_PiM:
   fixes M1 :: "'a measure" and M2 :: "'a measure"
   assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
-  shows "(M1 \<Otimes>\<^sub>M M2) = distr (Pi\<^sub>M UNIV (bool_case M1 M2)) (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. (x True, x False))"
+  shows "(M1 \<Otimes>\<^sub>M M2) = distr (Pi\<^sub>M UNIV (case_bool M1 M2)) (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. (x True, x False))"
     (is "?P = ?D")
 proof (rule pair_measure_eqI[OF assms])
-  interpret B: product_sigma_finite "bool_case M1 M2"
+  interpret B: product_sigma_finite "case_bool M1 M2"
     unfolding product_sigma_finite_def using assms by (auto split: bool.split)
-  let ?B = "Pi\<^sub>M UNIV (bool_case M1 M2)"
+  let ?B = "Pi\<^sub>M UNIV (case_bool M1 M2)"
 
   have [simp]: "fst \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x True)" "snd \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x False)"
     by auto
   fix A B assume A: "A \<in> sets M1" and B: "B \<in> sets M2"
-  have "emeasure M1 A * emeasure M2 B = (\<Prod> i\<in>UNIV. emeasure (bool_case M1 M2 i) (bool_case A B i))"
+  have "emeasure M1 A * emeasure M2 B = (\<Prod> i\<in>UNIV. emeasure (case_bool M1 M2 i) (case_bool A B i))"
     by (simp add: UNIV_bool ac_simps)
-  also have "\<dots> = emeasure ?B (Pi\<^sub>E UNIV (bool_case A B))"
+  also have "\<dots> = emeasure ?B (Pi\<^sub>E UNIV (case_bool A B))"
     using A B by (subst B.emeasure_PiM) (auto split: bool.split)
-  also have "Pi\<^sub>E UNIV (bool_case A B) = (\<lambda>x. (x True, x False)) -` (A \<times> B) \<inter> space ?B"
+  also have "Pi\<^sub>E UNIV (case_bool A B) = (\<lambda>x. (x True, x False)) -` (A \<times> B) \<inter> space ?B"
     using A[THEN sets.sets_into_space] B[THEN sets.sets_into_space]
     by (auto simp: PiE_iff all_bool_eq space_PiM split: bool.split)
   finally show "emeasure M1 A * emeasure M2 B = emeasure ?D (A \<times> B)"
     using A B
-      measurable_component_singleton[of True UNIV "bool_case M1 M2"]
-      measurable_component_singleton[of False UNIV "bool_case M1 M2"]
+      measurable_component_singleton[of True UNIV "case_bool M1 M2"]
+      measurable_component_singleton[of False UNIV "case_bool M1 M2"]
     by (subst emeasure_distr) (auto simp: measurable_pair_iff)
 qed simp
 
--- a/src/HOL/Probability/Independent_Family.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Probability/Independent_Family.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -13,7 +13,7 @@
     (\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> (\<forall>A\<in>Pi J F. prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))))"
 
 definition (in prob_space)
-  "indep_set A B \<longleftrightarrow> indep_sets (bool_case A B) UNIV"
+  "indep_set A B \<longleftrightarrow> indep_sets (case_bool A B) UNIV"
 
 definition (in prob_space)
   indep_events_def_alt: "indep_events A I \<longleftrightarrow> indep_sets (\<lambda>i. {A i}) I"
@@ -28,7 +28,7 @@
   done
 
 definition (in prob_space)
-  "indep_event A B \<longleftrightarrow> indep_events (bool_case A B) UNIV"
+  "indep_event A B \<longleftrightarrow> indep_events (case_bool A B) UNIV"
 
 lemma (in prob_space) indep_sets_cong:
   "I = J \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> F i = G i) \<Longrightarrow> indep_sets F I \<longleftrightarrow> indep_sets G J"
@@ -104,7 +104,7 @@
 lemma (in prob_space) indep_setD:
   assumes indep: "indep_set A B" and ev: "a \<in> A" "b \<in> B"
   shows "prob (a \<inter> b) = prob a * prob b"
-  using indep[unfolded indep_set_def, THEN indep_setsD, of UNIV "bool_case a b"] ev
+  using indep[unfolded indep_set_def, THEN indep_setsD, of UNIV "case_bool a b"] ev
   by (simp add: ac_simps UNIV_bool)
 
 lemma (in prob_space)
@@ -312,7 +312,7 @@
     indep_sets (\<lambda>i. { X i -` A \<inter> space M | A. A \<in> sets (M' i)}) I"
 
 definition (in prob_space)
-  "indep_var Ma A Mb B \<longleftrightarrow> indep_vars (bool_case Ma Mb) (bool_case A B) UNIV"
+  "indep_var Ma A Mb B \<longleftrightarrow> indep_vars (case_bool Ma Mb) (case_bool A B) UNIV"
 
 lemma (in prob_space) indep_vars_def:
   "indep_vars M' X I \<longleftrightarrow>
@@ -340,16 +340,16 @@
   "indep_set A B \<longleftrightarrow> A \<subseteq> events \<and> B \<subseteq> events \<and> (\<forall>a\<in>A. \<forall>b\<in>B. prob (a \<inter> b) = prob a * prob b)"
   unfolding indep_set_def
 proof (intro iffI ballI conjI)
-  assume indep: "indep_sets (bool_case A B) UNIV"
+  assume indep: "indep_sets (case_bool A B) UNIV"
   { fix a b assume "a \<in> A" "b \<in> B"
-    with indep_setsD[OF indep, of UNIV "bool_case a b"]
+    with indep_setsD[OF indep, of UNIV "case_bool a b"]
     show "prob (a \<inter> b) = prob a * prob b"
       unfolding UNIV_bool by (simp add: ac_simps) }
   from indep show "A \<subseteq> events" "B \<subseteq> events"
     unfolding indep_sets_def UNIV_bool by auto
 next
   assume *: "A \<subseteq> events \<and> B \<subseteq> events \<and> (\<forall>a\<in>A. \<forall>b\<in>B. prob (a \<inter> b) = prob a * prob b)"
-  show "indep_sets (bool_case A B) UNIV"
+  show "indep_sets (case_bool A B) UNIV"
   proof (rule indep_setsI)
     fix i show "(case i of True \<Rightarrow> A | False \<Rightarrow> B) \<subseteq> events"
       using * by (auto split: bool.split)
@@ -369,7 +369,7 @@
 proof -
   have "indep_sets (\<lambda>i. sigma_sets (space M) (case i of True \<Rightarrow> A | False \<Rightarrow> B)) UNIV"
   proof (rule indep_sets_sigma)
-    show "indep_sets (bool_case A B) UNIV"
+    show "indep_sets (case_bool A B) UNIV"
       by (rule `indep_set A B`[unfolded indep_set_def])
     fix i show "Int_stable (case i of True \<Rightarrow> A | False \<Rightarrow> B)"
       using A B by (cases i) auto
@@ -572,19 +572,19 @@
   qed
 
   { fix n
-    have "indep_sets (\<lambda>b. sigma_sets (space M) (\<Union>m\<in>bool_case {..n} {Suc n..} b. A m)) UNIV"
+    have "indep_sets (\<lambda>b. sigma_sets (space M) (\<Union>m\<in>case_bool {..n} {Suc n..} b. A m)) UNIV"
     proof (rule indep_sets_collect_sigma)
       have *: "(\<Union>b. case b of True \<Rightarrow> {..n} | False \<Rightarrow> {Suc n..}) = UNIV" (is "?U = _")
         by (simp split: bool.split add: set_eq_iff) (metis not_less_eq_eq)
       with indep show "indep_sets A ?U" by simp
-      show "disjoint_family (bool_case {..n} {Suc n..})"
+      show "disjoint_family (case_bool {..n} {Suc n..})"
         unfolding disjoint_family_on_def by (auto split: bool.split)
       fix m
       show "Int_stable (A m)"
         unfolding Int_stable_def using A.Int by auto
     qed
-    also have "(\<lambda>b. sigma_sets (space M) (\<Union>m\<in>bool_case {..n} {Suc n..} b. A m)) =
-      bool_case (sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) (sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m))"
+    also have "(\<lambda>b. sigma_sets (space M) (\<Union>m\<in>case_bool {..n} {Suc n..} b. A m)) =
+      case_bool (sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) (sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m))"
       by (auto intro!: ext split: bool.split)
     finally have indep: "indep_set (sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) (sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m))"
       unfolding indep_set_def by simp
@@ -923,9 +923,9 @@
     prob (A -` Xa \<inter> space M) * prob (B -` Xb \<inter> space M)"
 proof -
   have "prob ((\<lambda>x. (A x, B x)) -` (Xa \<times> Xb) \<inter> space M) =
-    prob (\<Inter>i\<in>UNIV. (bool_case A B i -` bool_case Xa Xb i \<inter> space M))"
+    prob (\<Inter>i\<in>UNIV. (case_bool A B i -` case_bool Xa Xb i \<inter> space M))"
     by (auto intro!: arg_cong[where f=prob] simp: UNIV_bool)
-  also have "\<dots> = (\<Prod>i\<in>UNIV. prob (bool_case A B i -` bool_case Xa Xb i \<inter> space M))"
+  also have "\<dots> = (\<Prod>i\<in>UNIV. prob (case_bool A B i -` case_bool Xa Xb i \<inter> space M))"
     using indep unfolding indep_var_def
     by (rule indep_varsD) (auto split: bool.split intro: sets)
   also have "\<dots> = prob (A -` Xa \<inter> space M) * prob (B -` Xb \<inter> space M)"
@@ -938,7 +938,7 @@
   shows indep_var_rv1: "random_variable S X"
     and indep_var_rv2: "random_variable T Y"
 proof -
-  have "\<forall>i\<in>UNIV. random_variable (bool_case S T i) (bool_case X Y i)"
+  have "\<forall>i\<in>UNIV. random_variable (case_bool S T i) (case_bool X Y i)"
     using assms unfolding indep_var_def indep_vars_def by auto
   then show "random_variable S X" "random_variable T Y"
     unfolding UNIV_bool by auto
--- a/src/HOL/Probability/Infinite_Product_Measure.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Probability/Infinite_Product_Measure.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -190,13 +190,13 @@
       let ?P =
         "\<lambda>k wk w. w \<in> space (Pi\<^sub>M (J (Suc k)) M) \<and> restrict w (J k) = wk \<and>
           (\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> ?q (Suc k) n w)"
-      def w \<equiv> "nat_rec w0 (\<lambda>k wk. Eps (?P k wk))"
+      def w \<equiv> "rec_nat w0 (\<lambda>k wk. Eps (?P k wk))"
 
       { fix k have w: "w k \<in> space (Pi\<^sub>M (J k) M) \<and>
           (\<forall>n. ?a / 2 ^ (k + 1) \<le> ?q k n (w k)) \<and> (k \<noteq> 0 \<longrightarrow> restrict (w k) (J (k - 1)) = w (k - 1))"
         proof (induct k)
           case 0 with w0 show ?case
-            unfolding w_def nat_rec_0 by auto
+            unfolding w_def nat.recs(1) by auto
         next
           case (Suc k)
           then have wk: "w k \<in> space (Pi\<^sub>M (J k) M)" by auto
@@ -241,7 +241,7 @@
                  (auto split: split_merge intro!: extensional_merge_sub ext simp: space_PiM PiE_iff)
           qed
           then have "?P k (w k) (w (Suc k))"
-            unfolding w_def nat_rec_Suc unfolding w_def[symmetric]
+            unfolding w_def nat.recs(2) unfolding w_def[symmetric]
             by (rule someI_ex)
           then show ?case by auto
         qed
@@ -464,11 +464,11 @@
   show "(\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)) \<rightarrow> (UNIV \<rightarrow>\<^sub>E space M)"
     by (auto simp: space_pair_measure space_PiM PiE_iff split: split_comb_seq)
   fix j :: nat and A assume A: "A \<in> sets M"
-  then have *: "{\<omega> \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)). prod_case (comb_seq i) \<omega> j \<in> A} =
+  then have *: "{\<omega> \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)). case_prod (comb_seq i) \<omega> j \<in> A} =
     (if j < i then {\<omega> \<in> space (\<Pi>\<^sub>M i\<in>UNIV. M). \<omega> j \<in> A} \<times> space (\<Pi>\<^sub>M i\<in>UNIV. M)
               else space (\<Pi>\<^sub>M i\<in>UNIV. M) \<times> {\<omega> \<in> space (\<Pi>\<^sub>M i\<in>UNIV. M). \<omega> (j - i) \<in> A})"
     by (auto simp: space_PiM space_pair_measure comb_seq_def dest: sets.sets_into_space)
-  show "{\<omega> \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)). prod_case (comb_seq i) \<omega> j \<in> A} \<in> sets ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M))"
+  show "{\<omega> \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)). case_prod (comb_seq i) \<omega> j \<in> A} \<in> sets ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M))"
     unfolding * by (auto simp: A intro!: sets_Collect_single)
 qed
 
@@ -480,10 +480,10 @@
 lemma comb_seq_0: "comb_seq 0 \<omega> \<omega>' = \<omega>'"
   by (auto simp add: comb_seq_def)
 
-lemma comb_seq_Suc: "comb_seq (Suc n) \<omega> \<omega>' = comb_seq n \<omega> (nat_case (\<omega> n) \<omega>')"
+lemma comb_seq_Suc: "comb_seq (Suc n) \<omega> \<omega>' = comb_seq n \<omega> (case_nat (\<omega> n) \<omega>')"
   by (auto simp add: comb_seq_def not_less less_Suc_eq le_imp_diff_is_add intro!: ext split: nat.split)
 
-lemma comb_seq_Suc_0[simp]: "comb_seq (Suc 0) \<omega> = nat_case (\<omega> 0)"
+lemma comb_seq_Suc_0[simp]: "comb_seq (Suc 0) \<omega> = case_nat (\<omega> 0)"
   by (intro ext) (simp add: comb_seq_Suc comb_seq_0)
 
 lemma comb_seq_less: "i < n \<Longrightarrow> comb_seq n \<omega> \<omega>' i = \<omega> i"
@@ -492,11 +492,11 @@
 lemma comb_seq_add: "comb_seq n \<omega> \<omega>' (i + n) = \<omega>' i"
   by (auto split: nat.split split_comb_seq)
 
-lemma nat_case_comb_seq: "nat_case s' (comb_seq n \<omega> \<omega>') (i + n) = nat_case (nat_case s' \<omega> n) \<omega>' i"
+lemma case_nat_comb_seq: "case_nat s' (comb_seq n \<omega> \<omega>') (i + n) = case_nat (case_nat s' \<omega> n) \<omega>' i"
   by (auto split: nat.split split_comb_seq)
 
-lemma nat_case_comb_seq':
-  "nat_case s (comb_seq i \<omega> \<omega>') = comb_seq (Suc i) (nat_case s \<omega>) \<omega>'"
+lemma case_nat_comb_seq':
+  "case_nat s (comb_seq i \<omega> \<omega>') = comb_seq (Suc i) (case_nat s \<omega>) \<omega>'"
   by (auto split: split_comb_seq nat.split)
 
 locale sequence_space = product_prob_space "\<lambda>i. M" "UNIV :: nat set" for M
@@ -570,7 +570,7 @@
 qed simp_all
 
 lemma PiM_iter:
-  "distr (M \<Otimes>\<^sub>M S) S (\<lambda>(s, \<omega>). nat_case s \<omega>) = S" (is "?D = _")
+  "distr (M \<Otimes>\<^sub>M S) S (\<lambda>(s, \<omega>). case_nat s \<omega>) = S" (is "?D = _")
 proof (rule PiM_eq)
   let ?I = "UNIV::nat set" and ?M = "\<lambda>n. M"
   let "distr _ _ ?f" = "?D"
--- a/src/HOL/Probability/Radon_Nikodym.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Probability/Radon_Nikodym.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -244,13 +244,13 @@
       by (simp add: measure_restricted sets_eq sets.Int) (metis inf_absorb2)
     hence "\<exists>A. ?P A S n" .. }
   note Ex_P = this
-  def A \<equiv> "nat_rec (space M) (\<lambda>n A. SOME B. ?P B A n)"
+  def A \<equiv> "rec_nat (space M) (\<lambda>n A. SOME B. ?P B A n)"
   have A_Suc: "\<And>n. A (Suc n) = (SOME B. ?P B (A n) n)" by (simp add: A_def)
   have A_0[simp]: "A 0 = space M" unfolding A_def by simp
   { fix i have "A i \<in> sets M" unfolding A_def
     proof (induct i)
       case (Suc i)
-      from Ex_P[OF this, of i] show ?case unfolding nat_rec_Suc
+      from Ex_P[OF this, of i] show ?case unfolding nat.recs(2)
         by (rule someI2_ex) simp
     qed simp }
   note A_in_sets = this
@@ -281,7 +281,7 @@
       from ex_inverse_of_nat_Suc_less[OF this]
       obtain n where *: "?d B < - 1 / real (Suc n)"
         by (auto simp: real_eq_of_nat inverse_eq_divide field_simps)
-      have "B \<subseteq> A (Suc n)" using B by (auto simp del: nat_rec_Suc)
+      have "B \<subseteq> A (Suc n)" using B by (auto simp del: nat.recs(2))
       from epsilon[OF B(1) this] *
       show False by auto
     qed
--- a/src/HOL/Product_Type.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Product_Type.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -12,8 +12,27 @@
 
 subsection {* @{typ bool} is a datatype *}
 
+wrap_free_constructors [True, False] case_bool [=]
+by auto
+
+-- {* Avoid name clashes by prefixing the output of @{text rep_datatype} with @{text old}. *}
+setup {* Sign.mandatory_path "old" *}
+
 rep_datatype True False by (auto intro: bool_induct)
 
+setup {* Sign.parent_path *}
+
+-- {* But erase the prefix for properties that are not generated by @{text wrap_free_constructors}. *}
+setup {* Sign.mandatory_path "bool" *}
+
+lemmas induct = old.bool.induct
+lemmas inducts = old.bool.inducts
+lemmas recs = old.bool.recs
+lemmas cases = bool.case
+lemmas simps = bool.distinct bool.case old.bool.recs
+
+setup {* Sign.parent_path *}
+
 declare case_split [cases type: bool]
   -- "prefer plain propositional version"
 
@@ -61,8 +80,27 @@
     else SOME (mk_meta_eq @{thm unit_eq})
 *}
 
+wrap_free_constructors ["()"] case_unit
+by auto
+
+-- {* Avoid name clashes by prefixing the output of @{text rep_datatype} with @{text old}. *}
+setup {* Sign.mandatory_path "old" *}
+
 rep_datatype "()" by simp
 
+setup {* Sign.parent_path *}
+
+-- {* But erase the prefix for properties that are not generated by @{text wrap_free_constructors}. *}
+setup {* Sign.mandatory_path "unit" *}
+
+lemmas induct = old.unit.induct
+lemmas inducts = old.unit.inducts
+lemmas recs = old.unit.recs
+lemmas cases = unit.case
+lemmas simps = unit.case old.unit.recs
+
+setup {* Sign.parent_path *}
+
 lemma unit_all_eq1: "(!!x::unit. PROP P x) == PROP P ()"
   by simp
 
@@ -139,10 +177,14 @@
 definition Pair :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<times> 'b" where
   "Pair a b = Abs_prod (Pair_Rep a b)"
 
-rep_datatype Pair proof -
-  fix P :: "'a \<times> 'b \<Rightarrow> bool" and p
-  assume "\<And>a b. P (Pair a b)"
-  then show "P p" by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def)
+lemma prod_cases: "(\<And>a b. P (Pair a b)) \<Longrightarrow> P p"
+  by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def)
+
+wrap_free_constructors [Pair] case_prod [] [[fst, snd]]
+proof -
+  fix P :: bool and p :: "'a \<times> 'b"
+  show "(\<And>x1 x2. p = Pair x1 x2 \<Longrightarrow> P) \<Longrightarrow> P"
+    by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def)
 next
   fix a c :: 'a and b d :: 'b
   have "Pair_Rep a b = Pair_Rep c d \<longleftrightarrow> a = c \<and> b = d"
@@ -153,15 +195,36 @@
     by (simp add: Pair_def Abs_prod_inject)
 qed
 
-declare prod.simps(2) [nitpick_simp del]
+-- {* Avoid name clashes by prefixing the output of @{text rep_datatype} with @{text old}. *}
+setup {* Sign.mandatory_path "old" *}
+
+rep_datatype Pair
+by (erule prod_cases) (rule prod.inject)
+
+setup {* Sign.parent_path *}
+
+-- {* But erase the prefix for properties that are not generated by @{text wrap_free_constructors}. *}
+setup {* Sign.mandatory_path "prod" *}
 
+declare
+  old.prod.inject[iff del]
+
+lemmas induct = old.prod.induct
+lemmas inducts = old.prod.inducts
+lemmas recs = old.prod.recs
+lemmas cases = prod.case
+lemmas simps = prod.inject prod.case old.prod.recs
+
+setup {* Sign.parent_path *}
+
+declare prod.case [nitpick_simp del]
 declare prod.weak_case_cong [cong del]
 
 
 subsubsection {* Tuple syntax *}
 
 abbreviation (input) split :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c" where
-  "split \<equiv> prod_case"
+  "split \<equiv> case_prod"
 
 text {*
   Patterns -- extends pre-defined type @{typ pttrn} used in
@@ -183,8 +246,8 @@
   "_pattern x y" => "CONST Pair x y"
   "_patterns x y" => "CONST Pair x y"
   "_tuple x (_tuple_args y z)" == "_tuple x (_tuple_arg (_tuple y z))"
-  "%(x, y, zs). b" == "CONST prod_case (%x (y, zs). b)"
-  "%(x, y). b" == "CONST prod_case (%x y. b)"
+  "%(x, y, zs). b" == "CONST case_prod (%x (y, zs). b)"
+  "%(x, y). b" == "CONST case_prod (%x y. b)"
   "_abs (CONST Pair x y) t" => "%(x, y). t"
   -- {* The last rule accommodates tuples in `case C ... (x,y) ... => ...'
      The (x,y) is parsed as `Pair x y' because it is logic, not pttrn *}
@@ -202,7 +265,7 @@
             Syntax.const @{syntax_const "_abs"} $
               (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t''
           end
-      | split_tr' [Abs (x, T, (s as Const (@{const_syntax prod_case}, _) $ t))] =
+      | split_tr' [Abs (x, T, (s as Const (@{const_syntax case_prod}, _) $ t))] =
           (* split (%x. (split (%y z. t))) => %(x,y,z). t *)
           let
             val Const (@{syntax_const "_abs"}, _) $
@@ -213,7 +276,7 @@
               (Syntax.const @{syntax_const "_pattern"} $ x' $
                 (Syntax.const @{syntax_const "_patterns"} $ y $ z)) $ t''
           end
-      | split_tr' [Const (@{const_syntax prod_case}, _) $ t] =
+      | split_tr' [Const (@{const_syntax case_prod}, _) $ t] =
           (* split (split (%x y z. t)) => %((x, y), z). t *)
           split_tr' [(split_tr' [t])] (* inner split_tr' creates next pattern *)
       | split_tr' [Const (@{syntax_const "_abs"}, _) $ x_y $ Abs abs] =
@@ -223,7 +286,7 @@
               (Syntax.const @{syntax_const "_pattern"} $ x_y $ z) $ t
           end
       | split_tr' _ = raise Match;
-  in [(@{const_syntax prod_case}, K split_tr')] end
+  in [(@{const_syntax case_prod}, K split_tr')] end
 *}
 
 (* print "split f" as "\<lambda>(x,y). f x y" and "split (\<lambda>x. f x)" as "\<lambda>(x,y). f x y" *) 
@@ -232,7 +295,7 @@
     fun split_guess_names_tr' T [Abs (x, _, Abs _)] = raise Match
       | split_guess_names_tr' T [Abs (x, xT, t)] =
           (case (head_of t) of
-            Const (@{const_syntax prod_case}, _) => raise Match
+            Const (@{const_syntax case_prod}, _) => raise Match
           | _ =>
             let 
               val (_ :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match;
@@ -244,7 +307,7 @@
             end)
       | split_guess_names_tr' T [t] =
           (case head_of t of
-            Const (@{const_syntax prod_case}, _) => raise Match
+            Const (@{const_syntax case_prod}, _) => raise Match
           | _ =>
             let
               val (xT :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match;
@@ -256,10 +319,10 @@
                 (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t''
             end)
       | split_guess_names_tr' _ _ = raise Match;
-  in [(@{const_syntax prod_case}, K split_guess_names_tr')] end
+  in [(@{const_syntax case_prod}, K split_guess_names_tr')] end
 *}
 
-(* Force eta-contraction for terms of the form "Q A (%p. prod_case P p)"
+(* Force eta-contraction for terms of the form "Q A (%p. case_prod P p)"
    where Q is some bounded quantifier or set operator.
    Reason: the above prints as "Q p : A. case p of (x,y) \<Rightarrow> P x y"
    whereas we want "Q (x,y):A. P x y".
@@ -269,7 +332,7 @@
   let
     fun contract Q tr ctxt ts =
       (case ts of
-        [A, Abs (_, _, (s as Const (@{const_syntax prod_case},_) $ t) $ Bound 0)] =>
+        [A, Abs (_, _, (s as Const (@{const_syntax case_prod},_) $ t) $ Bound 0)] =>
           if Term.is_dependent t then tr ctxt ts
           else Syntax.const Q $ A $ s
       | _ => tr ctxt ts);
@@ -312,23 +375,11 @@
 lemma surj_pair [simp]: "EX x y. p = (x, y)"
   by (cases p) simp
 
-definition fst :: "'a \<times> 'b \<Rightarrow> 'a" where
-  "fst p = (case p of (a, b) \<Rightarrow> a)"
-
-definition snd :: "'a \<times> 'b \<Rightarrow> 'b" where
-  "snd p = (case p of (a, b) \<Rightarrow> b)"
-
-lemma fst_conv [simp, code]: "fst (a, b) = a"
-  unfolding fst_def by simp
-
-lemma snd_conv [simp, code]: "snd (a, b) = b"
-  unfolding snd_def by simp
-
 code_printing
   constant fst \<rightharpoonup> (Haskell) "fst"
 | constant snd \<rightharpoonup> (Haskell) "snd"
 
-lemma prod_case_unfold [nitpick_unfold]: "prod_case = (%c p. c (fst p) (snd p))"
+lemma case_prod_unfold [nitpick_unfold]: "case_prod = (%c p. c (fst p) (snd p))"
   by (simp add: fun_eq_iff split: prod.split)
 
 lemma fst_eqD: "fst (x, y) = a ==> x = a"
@@ -337,10 +388,7 @@
 lemma snd_eqD: "snd (x, y) = a ==> y = a"
   by simp
 
-lemma pair_collapse [simp]: "(fst p, snd p) = p"
-  by (cases p) simp
-
-lemmas surjective_pairing = pair_collapse [symmetric]
+lemmas surjective_pairing = prod.collapse [symmetric]
 
 lemma prod_eq_iff: "s = t \<longleftrightarrow> fst s = fst t \<and> snd s = snd t"
   by (cases s, cases t) simp
@@ -371,7 +419,7 @@
   by (cases p) simp
 
 lemma The_split: "The (split P) = (THE xy. P (fst xy) (snd xy))"
-  by (simp add: prod_case_unfold)
+  by (simp add: case_prod_unfold)
 
 lemma split_weak_cong: "p = q \<Longrightarrow> split c p = split c q"
   -- {* Prevents simplification of @{term c}: much faster *}
@@ -463,7 +511,7 @@
     | no_args k i (Bound m) = m < k orelse m > k + i
     | no_args _ _ _ = true;
   fun split_pat tp i (Abs  (_, _, t)) = if tp 0 i t then SOME (i, t) else NONE
-    | split_pat tp i (Const (@{const_name prod_case}, _) $ Abs (_, _, t)) = split_pat tp (i + 1) t
+    | split_pat tp i (Const (@{const_name case_prod}, _) $ Abs (_, _, t)) = split_pat tp (i + 1) t
     | split_pat tp i _ = NONE;
   fun metaeq ctxt lhs rhs = mk_meta_eq (Goal.prove ctxt [] []
         (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)))
@@ -481,12 +529,12 @@
         else (subst arg k i t $ subst arg k i u)
     | subst arg k i t = t;
 in
-  fun beta_proc ctxt (s as Const (@{const_name prod_case}, _) $ Abs (_, _, t) $ arg) =
+  fun beta_proc ctxt (s as Const (@{const_name case_prod}, _) $ Abs (_, _, t) $ arg) =
         (case split_pat beta_term_pat 1 t of
           SOME (i, f) => SOME (metaeq ctxt s (subst arg 0 i f))
         | NONE => NONE)
     | beta_proc _ _ = NONE;
-  fun eta_proc ctxt (s as Const (@{const_name prod_case}, _) $ Abs (_, _, t)) =
+  fun eta_proc ctxt (s as Const (@{const_name case_prod}, _) $ Abs (_, _, t)) =
         (case split_pat eta_term_pat 1 t of
           SOME (_, ft) => SOME (metaeq ctxt s (let val (f $ arg) = ft in f end))
         | NONE => NONE)
@@ -563,14 +611,14 @@
   assumes major: "z \<in> split c p"
     and cases: "\<And>x y. p = (x, y) \<Longrightarrow> z \<in> c x y \<Longrightarrow> Q"
   shows Q
-  by (rule major [unfolded prod_case_unfold] cases surjective_pairing)+
+  by (rule major [unfolded case_prod_unfold] cases surjective_pairing)+
 
 declare mem_splitI2 [intro!] mem_splitI [intro!] splitI2' [intro!] splitI2 [intro!] splitI [intro!]
 declare mem_splitE [elim!] splitE' [elim!] splitE [elim!]
 
 ML {*
 local (* filtering with exists_p_split is an essential optimization *)
-  fun exists_p_split (Const (@{const_name prod_case},_) $ _ $ (Const (@{const_name Pair},_)$_$_)) = true
+  fun exists_p_split (Const (@{const_name case_prod},_) $ _ $ (Const (@{const_name Pair},_)$_$_)) = true
     | exists_p_split (t $ u) = exists_p_split t orelse exists_p_split u
     | exists_p_split (Abs (_, _, t)) = exists_p_split t
     | exists_p_split _ = false;
@@ -630,24 +678,24 @@
   Setup of internal @{text split_rule}.
 *}
 
-lemmas prod_caseI = prod.cases [THEN iffD2]
+lemmas case_prodI = prod.cases [THEN iffD2]
 
-lemma prod_caseI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> prod_case c p"
+lemma case_prodI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> case_prod c p"
   by (fact splitI2)
 
-lemma prod_caseI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> prod_case c p x"
+lemma case_prodI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> case_prod c p x"
   by (fact splitI2')
 
-lemma prod_caseE: "prod_case c p ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q"
+lemma case_prodE: "case_prod c p ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q"
   by (fact splitE)
 
-lemma prod_caseE': "prod_case c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q"
+lemma case_prodE': "case_prod c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q"
   by (fact splitE')
 
-declare prod_caseI [intro!]
+declare case_prodI [intro!]
 
-lemma prod_case_beta:
-  "prod_case f p = f (fst p) (snd p)"
+lemma case_prod_beta:
+  "case_prod f p = f (fst p) (snd p)"
   by (fact split_beta)
 
 lemma prod_cases3 [cases type]:
@@ -692,7 +740,7 @@
 
 lemma split_def:
   "split = (\<lambda>c p. c (fst p) (snd p))"
-  by (fact prod_case_unfold)
+  by (fact case_prod_unfold)
 
 definition internal_split :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c" where
   "internal_split == split"
@@ -739,13 +787,13 @@
 notation fcomp (infixl "\<circ>>" 60)
 
 definition scomp :: "('a \<Rightarrow> 'b \<times> 'c) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> 'd) \<Rightarrow> 'a \<Rightarrow> 'd" (infixl "\<circ>\<rightarrow>" 60) where
-  "f \<circ>\<rightarrow> g = (\<lambda>x. prod_case g (f x))"
+  "f \<circ>\<rightarrow> g = (\<lambda>x. case_prod g (f x))"
 
 lemma scomp_unfold: "scomp = (\<lambda>f g x. g (fst (f x)) (snd (f x)))"
-  by (simp add: fun_eq_iff scomp_def prod_case_unfold)
+  by (simp add: fun_eq_iff scomp_def case_prod_unfold)
 
-lemma scomp_apply [simp]: "(f \<circ>\<rightarrow> g) x = prod_case g (f x)"
-  by (simp add: scomp_unfold prod_case_unfold)
+lemma scomp_apply [simp]: "(f \<circ>\<rightarrow> g) x = case_prod g (f x)"
+  by (simp add: scomp_unfold case_prod_unfold)
 
 lemma Pair_scomp: "Pair x \<circ>\<rightarrow> f = f x"
   by (simp add: fun_eq_iff)
@@ -1179,9 +1227,10 @@
   by (fact prod.exhaust)
 
 lemmas Pair_eq = prod.inject
-
-lemmas split = split_conv  -- {* for backwards compatibility *}
-
+lemmas fst_conv = prod.sel(1)
+lemmas snd_conv = prod.sel(2)
+lemmas pair_collapse = prod.collapse
+lemmas split = split_conv
 lemmas Pair_fst_snd_eq = prod_eq_iff
 
 hide_const (open) prod
--- a/src/HOL/Proofs/Lambda/LambdaType.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Proofs/Lambda/LambdaType.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -30,12 +30,7 @@
   by (simp add: shift_def)
 
 lemma shift_commute [simp]: "e\<langle>i:U\<rangle>\<langle>0:T\<rangle> = e\<langle>0:T\<rangle>\<langle>Suc i:U\<rangle>"
-  apply (rule ext)
-  apply (case_tac x)
-   apply simp
-  apply (case_tac nat)
-   apply (simp_all add: shift_def)
-  done
+  by (rule ext) (simp_all add: shift_def split: nat.split)
 
 
 subsection {* Types and typing rules *}
@@ -157,6 +152,7 @@
     "e \<turnstile> t \<degree>\<degree> ts : T \<Longrightarrow> \<exists>Ts. e \<turnstile> t : Ts \<Rrightarrow> T \<and> e \<tturnstile> ts : Ts"
   apply (induct ts arbitrary: t T)
    apply simp
+  apply (rename_tac a b t T)
   apply atomize
   apply simp
   apply (erule_tac x = "t \<degree> a" in allE)
@@ -177,12 +173,14 @@
     "e \<turnstile> t : Ts \<Rrightarrow> T \<Longrightarrow> e \<tturnstile> ts : Ts \<Longrightarrow> e \<turnstile> t \<degree>\<degree> ts : T"
   apply (induct ts arbitrary: t T Ts)
    apply simp
+  apply (rename_tac a b t T Ts)
   apply atomize
   apply (case_tac Ts)
    apply simp
   apply simp
   apply (erule_tac x = "t \<degree> a" in allE)
   apply (erule_tac x = T in allE)
+  apply (rename_tac list)
   apply (erule_tac x = list in allE)
   apply (erule impE)
    apply (erule conjE)
@@ -225,6 +223,7 @@
   apply (erule var_app_type_eq)
   apply assumption
   apply simp
+  apply (rename_tac a b ts Ts U)
   apply atomize
   apply (case_tac U)
   apply (rule FalseE)
--- a/src/HOL/ROOT	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/ROOT	Wed Feb 12 10:59:25 2014 +0100
@@ -464,7 +464,7 @@
     Author:     Tobias Nipkow and Konrad Slind and Olaf Müller
     Copyright   1994--1996  TU Muenchen
 
-    The meta theory of I/O-Automata in HOL. This formalization has been
+    The meta-theory of I/O-Automata in HOL. This formalization has been
     significantly changed and extended, see HOLCF/IOA. There are also the
     proofs of two communication protocols which formerly have been here.
 
--- a/src/HOL/Relation.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Relation.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -54,7 +54,7 @@
 lemma lfp_induct2: -- {* Version of @{thm [source] lfp_induct} for binary relations *}
   "(a, b) \<in> lfp f \<Longrightarrow> mono f \<Longrightarrow>
     (\<And>a b. (a, b) \<in> f (lfp f \<inter> {(x, y). P x y}) \<Longrightarrow> P a b) \<Longrightarrow> P a b"
-  using lfp_induct_set [of "(a, b)" f "prod_case P"] by auto
+  using lfp_induct_set [of "(a, b)" f "case_prod P"] by auto
 
 
 subsubsection {* Conversions between set and predicate relations *}
@@ -113,7 +113,7 @@
 lemma INF_Int_eq [pred_set_conv]: "(\<Sqinter>i\<in>S. (\<lambda>x. x \<in> i)) = (\<lambda>x. x \<in> \<Inter>S)"
   by (simp add: fun_eq_iff)
 
-lemma Inf_INT_eq2 [pred_set_conv]: "\<Sqinter>S = (\<lambda>x y. (x, y) \<in> INTER (prod_case ` S) Collect)"
+lemma Inf_INT_eq2 [pred_set_conv]: "\<Sqinter>S = (\<lambda>x y. (x, y) \<in> INTER (case_prod ` S) Collect)"
   by (simp add: fun_eq_iff)
 
 lemma INF_Int_eq2 [pred_set_conv]: "(\<Sqinter>i\<in>S. (\<lambda>x y. (x, y) \<in> i)) = (\<lambda>x y. (x, y) \<in> \<Inter>S)"
@@ -125,7 +125,7 @@
 lemma SUP_Sup_eq [pred_set_conv]: "(\<Squnion>i\<in>S. (\<lambda>x. x \<in> i)) = (\<lambda>x. x \<in> \<Union>S)"
   by (simp add: fun_eq_iff)
 
-lemma Sup_SUP_eq2 [pred_set_conv]: "\<Squnion>S = (\<lambda>x y. (x, y) \<in> UNION (prod_case ` S) Collect)"
+lemma Sup_SUP_eq2 [pred_set_conv]: "\<Squnion>S = (\<lambda>x y. (x, y) \<in> UNION (case_prod ` S) Collect)"
   by (simp add: fun_eq_iff)
 
 lemma SUP_Sup_eq2 [pred_set_conv]: "(\<Squnion>i\<in>S. (\<lambda>x y. (x, y) \<in> i)) = (\<lambda>x y. (x, y) \<in> \<Union>S)"
--- a/src/HOL/SET_Protocol/Cardholder_Registration.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/SET_Protocol/Cardholder_Registration.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -671,6 +671,7 @@
 lemma N_fresh_not_KeyCryptNonce:
      "\<forall>C. DK \<noteq> priEK C ==> Nonce N \<notin> used evs --> ~ KeyCryptNonce DK N evs"
 apply (induct_tac "evs")
+apply (rename_tac [2] a evs')
 apply (case_tac [2] "a")
 apply (auto simp add: parts_insert2)
 done
--- a/src/HOL/SET_Protocol/Event_SET.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/SET_Protocol/Event_SET.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -159,6 +159,7 @@
 lemma Notes_imp_parts_subset_used [rule_format]:
      "Notes A X \<in> set evs --> parts {X} <= used evs"
 apply (induct_tac "evs")
+apply (rename_tac [2] a evs')
 apply (induct_tac [2] "a", auto)
 done
 
--- a/src/HOL/SET_Protocol/Message_SET.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/SET_Protocol/Message_SET.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -80,7 +80,7 @@
 
 
 definition nat_of_agent :: "agent => nat" where
-   "nat_of_agent == agent_case (curry prod_encode 0)
+   "nat_of_agent == case_agent (curry prod_encode 0)
                                (curry prod_encode 1)
                                (curry prod_encode 2)
                                (curry prod_encode 3)
--- a/src/HOL/SMT_Examples/SMT_Examples.certs	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/SMT_Examples/SMT_Examples.certs	Wed Feb 12 10:59:25 2014 +0100
@@ -101,40 +101,6 @@
 #45 := [and-elim #44]: #11
 [mp #45 #65]: false
 unsat
-050883983ebe99dc3b7f24a011b1724b1b2c4dd9 33 0
-#2 := false
-decl f1 :: S1
-#3 := f1
-decl f6 :: S1
-#14 := f6
-#15 := (= f6 f1)
-decl f5 :: S1
-#12 := f5
-#13 := (= f5 f1)
-#16 := (and #13 #15)
-decl f4 :: S1
-#9 := f4
-#10 := (= f4 f1)
-decl f3 :: S1
-#7 := f3
-#8 := (= f3 f1)
-#11 := (and #8 #10)
-#17 := (or #11 #16)
-#18 := (implies #17 #17)
-#19 := (not #18)
-#48 := (iff #19 false)
-#1 := true
-#43 := (not true)
-#46 := (iff #43 false)
-#47 := [rewrite]: #46
-#44 := (iff #19 #43)
-#41 := (iff #18 true)
-#42 := [rewrite]: #41
-#45 := [monotonicity #42]: #44
-#49 := [trans #45 #47]: #48
-#40 := [asserted]: #19
-[mp #40 #49]: false
-unsat
 79d9d246dd9d27e03e8f1ea895e790f3a4420bfd 55 0
 #2 := false
 decl f1 :: S1
@@ -191,6 +157,40 @@
 #42 := [asserted]: #21
 [mp #42 #72]: false
 unsat
+050883983ebe99dc3b7f24a011b1724b1b2c4dd9 33 0
+#2 := false
+decl f1 :: S1
+#3 := f1
+decl f6 :: S1
+#14 := f6
+#15 := (= f6 f1)
+decl f5 :: S1
+#12 := f5
+#13 := (= f5 f1)
+#16 := (and #13 #15)
+decl f4 :: S1
+#9 := f4
+#10 := (= f4 f1)
+decl f3 :: S1
+#7 := f3
+#8 := (= f3 f1)
+#11 := (and #8 #10)
+#17 := (or #11 #16)
+#18 := (implies #17 #17)
+#19 := (not #18)
+#48 := (iff #19 false)
+#1 := true
+#43 := (not true)
+#46 := (iff #43 false)
+#47 := [rewrite]: #46
+#44 := (iff #19 #43)
+#41 := (iff #18 true)
+#42 := [rewrite]: #41
+#45 := [monotonicity #42]: #44
+#49 := [trans #45 #47]: #48
+#40 := [asserted]: #19
+[mp #40 #49]: false
+unsat
 8575241c64c02491d277f6598ca57e576f5a6b45 60 0
 #2 := false
 decl f1 :: S1
@@ -546,6 +546,124 @@
 #83 := [and-elim #82]: #57
 [unit-resolution #83 #90]: false
 unsat
+a69a9e8c5e31ec6b9da4cf96f47b52cf6b9404d9 117 0
+#2 := false
+decl f3 :: (-> S3 S2 S1)
+#10 := (:var 0 S2)
+decl f4 :: (-> S4 S1 S3)
+decl f6 :: S1
+#16 := f6
+decl f5 :: S4
+#7 := f5
+#17 := (f4 f5 f6)
+#18 := (f3 #17 #10)
+#573 := (pattern #18)
+decl f1 :: S1
+#3 := f1
+#19 := (= #18 f1)
+#76 := (not #19)
+#574 := (forall (vars (?v0 S2)) (:pat #573) #76)
+decl f7 :: S2
+#21 := f7
+#22 := (f3 #17 f7)
+#23 := (= #22 f1)
+#150 := (= f6 f1)
+#151 := (iff #23 #150)
+#8 := (:var 1 S1)
+#9 := (f4 f5 #8)
+#11 := (f3 #9 #10)
+#566 := (pattern #11)
+#13 := (= #8 f1)
+#12 := (= #11 f1)
+#14 := (iff #12 #13)
+#567 := (forall (vars (?v0 S1) (?v1 S2)) (:pat #566) #14)
+#15 := (forall (vars (?v0 S1) (?v1 S2)) #14)
+#570 := (iff #15 #567)
+#568 := (iff #14 #14)
+#569 := [refl]: #568
+#571 := [quant-intro #569]: #570
+#62 := (~ #15 #15)
+#60 := (~ #14 #14)
+#61 := [refl]: #60
+#63 := [nnf-pos #61]: #62
+#46 := [asserted]: #15
+#53 := [mp~ #46 #63]: #15
+#572 := [mp #53 #571]: #567
+#152 := (not #567)
+#228 := (or #152 #151)
+#561 := [quant-inst #16 #21]: #228
+#237 := [unit-resolution #561 #572]: #151
+decl ?v0!0 :: S2
+#66 := ?v0!0
+#67 := (f3 #17 ?v0!0)
+#68 := (= #67 f1)
+#236 := (iff #68 #150)
+#238 := (or #152 #236)
+#229 := [quant-inst #16 #66]: #238
+#227 := [unit-resolution #229 #572]: #236
+#240 := (not #236)
+#199 := (or #240 #150)
+#55 := (not #23)
+#215 := [hypothesis]: #55
+#83 := (or #68 #23)
+#79 := (forall (vars (?v0 S2)) #76)
+#82 := (or #79 #55)
+#84 := (and #83 #82)
+#20 := (exists (vars (?v0 S2)) #19)
+#48 := (not #20)
+#49 := (iff #48 #23)
+#85 := (~ #49 #84)
+#57 := (~ #23 #23)
+#65 := [refl]: #57
+#64 := (~ #55 #55)
+#56 := [refl]: #64
+#80 := (~ #48 #79)
+#77 := (~ #76 #76)
+#78 := [refl]: #77
+#81 := [nnf-neg #78]: #80
+#73 := (not #48)
+#74 := (~ #73 #68)
+#69 := (~ #20 #68)
+#70 := [sk]: #69
+#75 := [nnf-neg #70]: #74
+#86 := [nnf-pos #75 #81 #56 #65]: #85
+#24 := (iff #20 #23)
+#25 := (not #24)
+#50 := (iff #25 #49)
+#51 := [rewrite]: #50
+#47 := [asserted]: #25
+#54 := [mp #47 #51]: #49
+#87 := [mp~ #54 #86]: #84
+#90 := [and-elim #87]: #83
+#557 := [unit-resolution #90 #215]: #68
+#243 := (not #68)
+#222 := (or #240 #243 #150)
+#558 := [def-axiom]: #222
+#541 := [unit-resolution #558 #557]: #199
+#203 := [unit-resolution #541 #227]: #150
+#241 := (not #150)
+#562 := (not #151)
+#204 := (or #562 #241)
+#563 := (or #562 #23 #241)
+#564 := [def-axiom]: #563
+#205 := [unit-resolution #564 #215]: #204
+#206 := [unit-resolution #205 #203 #237]: false
+#543 := [lemma #206]: #23
+#579 := (or #574 #55)
+#580 := (iff #82 #579)
+#577 := (iff #79 #574)
+#575 := (iff #76 #76)
+#576 := [refl]: #575
+#578 := [quant-intro #576]: #577
+#581 := [monotonicity #578]: #580
+#91 := [and-elim #87]: #82
+#582 := [mp #91 #581]: #579
+#242 := [unit-resolution #582 #543]: #574
+#555 := (not #574)
+#214 := (or #555 #55)
+#219 := [quant-inst #21]: #214
+[unit-resolution #219 #543 #242]: false
+unsat
 d97439af6f5bc7794ab403d0f6cc318d103016a1 1288 0
 #2 := false
 decl f1 :: S1
@@ -1835,124 +1953,6 @@
 #1532 := [unit-resolution #769 #1531]: #20
 [unit-resolution #606 #1532 #1528]: false
 unsat
-a69a9e8c5e31ec6b9da4cf96f47b52cf6b9404d9 117 0
-#2 := false
-decl f3 :: (-> S3 S2 S1)
-#10 := (:var 0 S2)
-decl f4 :: (-> S4 S1 S3)
-decl f6 :: S1
-#16 := f6
-decl f5 :: S4
-#7 := f5
-#17 := (f4 f5 f6)
-#18 := (f3 #17 #10)
-#573 := (pattern #18)
-decl f1 :: S1
-#3 := f1
-#19 := (= #18 f1)
-#76 := (not #19)
-#574 := (forall (vars (?v0 S2)) (:pat #573) #76)
-decl f7 :: S2
-#21 := f7
-#22 := (f3 #17 f7)
-#23 := (= #22 f1)
-#150 := (= f6 f1)
-#151 := (iff #23 #150)
-#8 := (:var 1 S1)
-#9 := (f4 f5 #8)
-#11 := (f3 #9 #10)
-#566 := (pattern #11)
-#13 := (= #8 f1)
-#12 := (= #11 f1)
-#14 := (iff #12 #13)
-#567 := (forall (vars (?v0 S1) (?v1 S2)) (:pat #566) #14)
-#15 := (forall (vars (?v0 S1) (?v1 S2)) #14)
-#570 := (iff #15 #567)
-#568 := (iff #14 #14)
-#569 := [refl]: #568
-#571 := [quant-intro #569]: #570
-#62 := (~ #15 #15)
-#60 := (~ #14 #14)
-#61 := [refl]: #60
-#63 := [nnf-pos #61]: #62
-#46 := [asserted]: #15
-#53 := [mp~ #46 #63]: #15
-#572 := [mp #53 #571]: #567
-#152 := (not #567)
-#228 := (or #152 #151)
-#561 := [quant-inst #16 #21]: #228
-#237 := [unit-resolution #561 #572]: #151
-decl ?v0!0 :: S2
-#66 := ?v0!0
-#67 := (f3 #17 ?v0!0)
-#68 := (= #67 f1)
-#236 := (iff #68 #150)
-#238 := (or #152 #236)
-#229 := [quant-inst #16 #66]: #238
-#227 := [unit-resolution #229 #572]: #236
-#240 := (not #236)
-#199 := (or #240 #150)
-#55 := (not #23)
-#215 := [hypothesis]: #55
-#83 := (or #68 #23)
-#79 := (forall (vars (?v0 S2)) #76)
-#82 := (or #79 #55)
-#84 := (and #83 #82)
-#20 := (exists (vars (?v0 S2)) #19)
-#48 := (not #20)
-#49 := (iff #48 #23)
-#85 := (~ #49 #84)
-#57 := (~ #23 #23)
-#65 := [refl]: #57
-#64 := (~ #55 #55)
-#56 := [refl]: #64
-#80 := (~ #48 #79)
-#77 := (~ #76 #76)
-#78 := [refl]: #77
-#81 := [nnf-neg #78]: #80
-#73 := (not #48)
-#74 := (~ #73 #68)
-#69 := (~ #20 #68)
-#70 := [sk]: #69
-#75 := [nnf-neg #70]: #74
-#86 := [nnf-pos #75 #81 #56 #65]: #85
-#24 := (iff #20 #23)
-#25 := (not #24)
-#50 := (iff #25 #49)
-#51 := [rewrite]: #50
-#47 := [asserted]: #25
-#54 := [mp #47 #51]: #49
-#87 := [mp~ #54 #86]: #84
-#90 := [and-elim #87]: #83
-#557 := [unit-resolution #90 #215]: #68
-#243 := (not #68)
-#222 := (or #240 #243 #150)
-#558 := [def-axiom]: #222
-#541 := [unit-resolution #558 #557]: #199
-#203 := [unit-resolution #541 #227]: #150
-#241 := (not #150)
-#562 := (not #151)
-#204 := (or #562 #241)
-#563 := (or #562 #23 #241)
-#564 := [def-axiom]: #563
-#205 := [unit-resolution #564 #215]: #204
-#206 := [unit-resolution #205 #203 #237]: false
-#543 := [lemma #206]: #23
-#579 := (or #574 #55)
-#580 := (iff #82 #579)
-#577 := (iff #79 #574)
-#575 := (iff #76 #76)
-#576 := [refl]: #575
-#578 := [quant-intro #576]: #577
-#581 := [monotonicity #578]: #580
-#91 := [and-elim #87]: #82
-#582 := [mp #91 #581]: #579
-#242 := [unit-resolution #582 #543]: #574
-#555 := (not #574)
-#214 := (or #555 #55)
-#219 := [quant-inst #21]: #214
-[unit-resolution #219 #543 #242]: false
-unsat
 fdf61e060f49731790f4d6c8f9b26c21349c60b3 117 0
 #2 := false
 decl f1 :: S1
@@ -2071,24 +2071,6 @@
 #603 := [unit-resolution #271 #618]: #602
 [unit-resolution #603 #601 #297]: false
 unsat
-0ce3a745d60cdbf0fe26b07c5e76de09d459dd25 17 0
-#2 := false
-#7 := 3::Int
-#8 := (= 3::Int 3::Int)
-#9 := (not #8)
-#38 := (iff #9 false)
-#1 := true
-#33 := (not true)
-#36 := (iff #33 false)
-#37 := [rewrite]: #36
-#34 := (iff #9 #33)
-#31 := (iff #8 true)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#39 := [trans #35 #37]: #38
-#30 := [asserted]: #9
-[mp #30 #39]: false
-unsat
 5c792581e65682628e5c59ca9f3f8801e6aeba72 61 0
 #2 := false
 decl f1 :: S1
@@ -2151,6 +2133,24 @@
 #136 := [quant-inst #7]: #221
 [unit-resolution #136 #556 #52]: false
 unsat
+0ce3a745d60cdbf0fe26b07c5e76de09d459dd25 17 0
+#2 := false
+#7 := 3::Int
+#8 := (= 3::Int 3::Int)
+#9 := (not #8)
+#38 := (iff #9 false)
+#1 := true
+#33 := (not true)
+#36 := (iff #33 false)
+#37 := [rewrite]: #36
+#34 := (iff #9 #33)
+#31 := (iff #8 true)
+#32 := [rewrite]: #31
+#35 := [monotonicity #32]: #34
+#39 := [trans #35 #37]: #38
+#30 := [asserted]: #9
+[mp #30 #39]: false
+unsat
 1532b1dde71eb42ca0a012bb62d9bbadf37fa326 17 0
 #2 := false
 #7 := 3::Real
@@ -7256,9 +7256,9 @@
 unsat
 c4f4c8220660d1979009b33a643f0927bee816b1 1 0
 unsat
-db6426d59fdd57da8ca5d11de399761d1f1443de 1 0
+e7ef76d73ccb9bc09d2b5368495a7a59d1bae3dc 1 0
 unsat
-e7ef76d73ccb9bc09d2b5368495a7a59d1bae3dc 1 0
+db6426d59fdd57da8ca5d11de399761d1f1443de 1 0
 unsat
 a2da5fa16f268876e3dcbc1874e34212d0a36218 54 0
 #2 := false
@@ -7799,6 +7799,70 @@
 #63 := [mp~ #61 #70]: #56
 [unit-resolution #63 #529]: false
 unsat
+f6f0c702e5caae5d1fc0a3e7862c44d261de6d47 63 0
+#2 := false
+#15 := 1::Int
+#12 := (:var 1 Int)
+#10 := 6::Int
+#11 := (- 6::Int)
+#13 := (* #11 #12)
+#8 := (:var 2 Int)
+#7 := 4::Int
+#9 := (* 4::Int #8)
+#14 := (+ #9 #13)
+#16 := (= #14 1::Int)
+#17 := (exists (vars (?v0 Int) (?v1 Int) (?v2 Int)) #16)
+#18 := (not #17)
+#19 := (not #18)
+#86 := (iff #19 false)
+#56 := (:var 0 Int)
+#41 := -6::Int
+#58 := (* -6::Int #56)
+#57 := (* 4::Int #12)
+#59 := (+ #57 #58)
+#60 := (= #59 1::Int)
+#61 := (exists (vars (?v0 Int) (?v1 Int)) #60)
+#84 := (iff #61 false)
+#77 := (exists (vars (?v0 Int) (?v1 Int)) false)
+#82 := (iff #77 false)
+#83 := [elim-unused]: #82
+#80 := (iff #61 #77)
+#78 := (iff #60 false)
+#79 := [rewrite]: #78
+#81 := [quant-intro #79]: #80
+#85 := [trans #81 #83]: #84
+#74 := (iff #19 #61)
+#66 := (not #61)
+#69 := (not #66)
+#72 := (iff #69 #61)
+#73 := [rewrite]: #72
+#70 := (iff #19 #69)
+#67 := (iff #18 #66)
+#64 := (iff #17 #61)
+#44 := (* -6::Int #12)
+#47 := (+ #9 #44)
+#50 := (= #47 1::Int)
+#53 := (exists (vars (?v0 Int) (?v1 Int) (?v2 Int)) #50)
+#62 := (iff #53 #61)
+#63 := [elim-unused]: #62
+#54 := (iff #17 #53)
+#51 := (iff #16 #50)
+#48 := (= #14 #47)
+#45 := (= #13 #44)
+#42 := (= #11 -6::Int)
+#43 := [rewrite]: #42
+#46 := [monotonicity #43]: #45
+#49 := [monotonicity #46]: #48
+#52 := [monotonicity #49]: #51
+#55 := [quant-intro #52]: #54
+#65 := [trans #55 #63]: #64
+#68 := [monotonicity #65]: #67
+#71 := [monotonicity #68]: #70
+#75 := [trans #71 #73]: #74
+#87 := [trans #75 #85]: #86
+#40 := [asserted]: #19
+[mp #40 #87]: false
+unsat
 252d255c564463d916bc68156eea8dbe7fb0be0a 165 0
 WARNING: failed to find a pattern for quantifier (quantifier id: k!10)
 #2 := false
@@ -7965,70 +8029,6 @@
 #563 := [unit-resolution #136 #574]: #62
 [unit-resolution #563 #570]: false
 unsat
-f6f0c702e5caae5d1fc0a3e7862c44d261de6d47 63 0
-#2 := false
-#15 := 1::Int
-#12 := (:var 1 Int)
-#10 := 6::Int
-#11 := (- 6::Int)
-#13 := (* #11 #12)
-#8 := (:var 2 Int)
-#7 := 4::Int
-#9 := (* 4::Int #8)
-#14 := (+ #9 #13)
-#16 := (= #14 1::Int)
-#17 := (exists (vars (?v0 Int) (?v1 Int) (?v2 Int)) #16)
-#18 := (not #17)
-#19 := (not #18)
-#86 := (iff #19 false)
-#56 := (:var 0 Int)
-#41 := -6::Int
-#58 := (* -6::Int #56)
-#57 := (* 4::Int #12)
-#59 := (+ #57 #58)
-#60 := (= #59 1::Int)
-#61 := (exists (vars (?v0 Int) (?v1 Int)) #60)
-#84 := (iff #61 false)
-#77 := (exists (vars (?v0 Int) (?v1 Int)) false)
-#82 := (iff #77 false)
-#83 := [elim-unused]: #82
-#80 := (iff #61 #77)
-#78 := (iff #60 false)
-#79 := [rewrite]: #78
-#81 := [quant-intro #79]: #80
-#85 := [trans #81 #83]: #84
-#74 := (iff #19 #61)
-#66 := (not #61)
-#69 := (not #66)
-#72 := (iff #69 #61)
-#73 := [rewrite]: #72
-#70 := (iff #19 #69)
-#67 := (iff #18 #66)
-#64 := (iff #17 #61)
-#44 := (* -6::Int #12)
-#47 := (+ #9 #44)
-#50 := (= #47 1::Int)
-#53 := (exists (vars (?v0 Int) (?v1 Int) (?v2 Int)) #50)
-#62 := (iff #53 #61)
-#63 := [elim-unused]: #62
-#54 := (iff #17 #53)
-#51 := (iff #16 #50)
-#48 := (= #14 #47)
-#45 := (= #13 #44)
-#42 := (= #11 -6::Int)
-#43 := [rewrite]: #42
-#46 := [monotonicity #43]: #45
-#49 := [monotonicity #46]: #48
-#52 := [monotonicity #49]: #51
-#55 := [quant-intro #52]: #54
-#65 := [trans #55 #63]: #64
-#68 := [monotonicity #65]: #67
-#71 := [monotonicity #68]: #70
-#75 := [trans #71 #73]: #74
-#87 := [trans #75 #85]: #86
-#40 := [asserted]: #19
-[mp #40 #87]: false
-unsat
 302156fb98e1f9b5657a3c89c418d5e1813f274a 101 0
 #2 := false
 #7 := 0::Int
@@ -10281,33 +10281,33 @@
 #55 := [not-or-elim #54]: #53
 [unit-resolution #55 #214]: false
 unsat
-22f5a208d6aa87f9794b1ab4d7ebb0a58f9ec89d 106 0
+86345bce2206ce27e174d4b1d6d3e0182564f8a1 106 0
 #2 := false
-decl f11 :: (-> S9 S7 S2)
-decl f16 :: S7
+decl f11 :: (-> S9 S5 S3)
+decl f16 :: S5
 #34 := f16
 decl f12 :: S9
 #25 := f12
 #39 := (f11 f12 f16)
-decl f3 :: (-> S4 S5 S2)
-decl f13 :: S5
+decl f6 :: (-> S6 S7 S3)
+decl f13 :: S7
 #29 := f13
-decl f4 :: S4
-#7 := f4
-#38 := (f3 f4 f13)
+decl f7 :: S6
+#14 := f7
+#38 := (f6 f7 f13)
 #40 := (= #38 #39)
-decl f8 :: (-> S3 S2 S7)
-decl f14 :: S2
+decl f5 :: (-> S2 S3 S5)
+decl f14 :: S3
 #30 := f14
-decl f15 :: S3
+decl f15 :: S2
 #31 := f15
-#35 := (f8 f15 f14)
-#236 := (f11 f12 #35)
-#233 := (= #236 #39)
-#573 := (= #39 #236)
+#35 := (f5 f15 f14)
+#165 := (f11 f12 #35)
+#233 := (= #165 #39)
+#573 := (= #39 #165)
 #36 := (= f16 #35)
-decl f5 :: (-> S2 S3 S5)
-#32 := (f5 f14 f15)
+decl f8 :: (-> S3 S2 S7)
+#32 := (f8 f14 f15)
 #33 := (= f13 #32)
 #37 := (and #33 #36)
 #68 := (not #37)
@@ -10325,17 +10325,17 @@
 #78 := [and-elim #75]: #36
 #579 := [monotonicity #78]: #573
 #570 := [symm #579]: #233
-#213 := (= #38 #236)
-#569 := (= f14 #236)
-#572 := (= #236 f14)
-#16 := (:var 0 S2)
-#15 := (:var 1 S3)
-#17 := (f8 #15 #16)
-#587 := (pattern #17)
-#26 := (f11 f12 #17)
-#27 := (= #26 #16)
-#600 := (forall (vars (?v0 S3) (?v1 S2)) (:pat #587) #27)
-#28 := (forall (vars (?v0 S3) (?v1 S2)) #27)
+#213 := (= #38 #165)
+#569 := (= f14 #165)
+#251 := (= #165 f14)
+#9 := (:var 0 S3)
+#8 := (:var 1 S2)
+#10 := (f5 #8 #9)
+#580 := (pattern #10)
+#26 := (f11 f12 #10)
+#27 := (= #26 #9)
+#600 := (forall (vars (?v0 S2) (?v1 S3)) (:pat #580) #27)
+#28 := (forall (vars (?v0 S2) (?v1 S3)) #27)
 #603 := (iff #28 #600)
 #601 := (iff #27 #27)
 #602 := [refl]: #601
@@ -10347,38 +10347,38 @@
 #66 := [asserted]: #28
 #109 := [mp~ #66 #89]: #28
 #605 := [mp #109 #604]: #600
-#242 := (not #600)
-#575 := (or #242 #572)
-#576 := [quant-inst #31 #30]: #575
-#568 := [unit-resolution #576 #605]: #572
+#256 := (not #600)
+#253 := (or #256 #251)
+#257 := [quant-inst #31 #30]: #253
+#568 := [unit-resolution #257 #605]: #251
 #228 := [symm #568]: #569
 #229 := (= #38 f14)
-#164 := (f3 f4 #32)
-#250 := (= #164 f14)
-#9 := (:var 0 S3)
-#8 := (:var 1 S2)
-#10 := (f5 #8 #9)
-#580 := (pattern #10)
-#11 := (f3 f4 #10)
-#12 := (= #11 #8)
-#581 := (forall (vars (?v0 S2) (?v1 S3)) (:pat #580) #12)
-#13 := (forall (vars (?v0 S2) (?v1 S3)) #12)
-#584 := (iff #13 #581)
-#582 := (iff #12 #12)
-#583 := [refl]: #582
-#585 := [quant-intro #583]: #584
-#100 := (~ #13 #13)
-#98 := (~ #12 #12)
-#99 := [refl]: #98
-#101 := [nnf-pos #99]: #100
-#63 := [asserted]: #13
-#82 := [mp~ #63 #101]: #13
-#586 := [mp #82 #585]: #581
-#166 := (not #581)
-#252 := (or #166 #250)
-#243 := [quant-inst #30 #31]: #252
-#241 := [unit-resolution #243 #586]: #250
-#577 := (= #38 #164)
+#254 := (f6 f7 #32)
+#255 := (= #254 f14)
+#16 := (:var 0 S2)
+#15 := (:var 1 S3)
+#17 := (f8 #15 #16)
+#587 := (pattern #17)
+#18 := (f6 f7 #17)
+#19 := (= #18 #15)
+#588 := (forall (vars (?v0 S3) (?v1 S2)) (:pat #587) #19)
+#20 := (forall (vars (?v0 S3) (?v1 S2)) #19)
+#591 := (iff #20 #588)
+#589 := (iff #19 #19)
+#590 := [refl]: #589
+#592 := [quant-intro #590]: #591
+#84 := (~ #20 #20)
+#83 := (~ #19 #19)
+#102 := [refl]: #83
+#85 := [nnf-pos #102]: #84
+#64 := [asserted]: #20
+#103 := [mp~ #64 #85]: #20
+#593 := [mp #103 #592]: #588
+#574 := (not #588)
+#230 := (or #574 #255)
+#361 := [quant-inst #30 #31]: #230
+#241 := [unit-resolution #361 #593]: #255
+#577 := (= #38 #254)
 #76 := [and-elim #75]: #33
 #578 := [monotonicity #76]: #577
 #571 := [trans #578 #241]: #229
--- a/src/HOL/SMT_Examples/SMT_Tests.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/SMT_Examples/SMT_Tests.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -669,13 +669,13 @@
   "tl [x, y, z] = [y, z]"
   "hd (tl [x, y, z]) = y"
   "tl (tl [x, y, z]) = [z]"
-  using hd.simps tl.simps(2) list.simps
+  using list.sel(1,3) list.simps
   by smt+
 
 lemma
   "fst (hd [(a, b)]) = a"
   "snd (hd [(a, b)]) = b"
-  using fst_conv snd_conv pair_collapse hd.simps tl.simps(2) list.simps
+  using fst_conv snd_conv pair_collapse list.sel(1,3) list.simps
   by smt+
 
 
@@ -808,14 +808,14 @@
   "tl [x, y, z] = [y, z]"
   "hd (tl [x, y, z]) = y"
   "tl (tl [x, y, z]) = [z]"
-  using hd.simps tl.simps(2)
+  using list.sel(1,3)
   using [[smt_datatypes, smt_oracle, z3_with_extensions]]
   by smt+
 
 lemma
   "fst (hd [(a, b)]) = a"
   "snd (hd [(a, b)]) = b"
-  using fst_conv snd_conv pair_collapse hd.simps tl.simps(2)
+  using fst_conv snd_conv pair_collapse list.sel(1,3)
   using [[smt_datatypes, smt_oracle, z3_with_extensions]]
   by smt+
 
--- a/src/HOL/String.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/String.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -286,13 +286,13 @@
 
 code_datatype Char -- {* drop case certificate for char *}
 
-lemma char_case_code [code]:
-  "char_case f c = (let n = nat_of_char c in f (nibble_of_nat (n div 16)) (nibble_of_nat n))"
+lemma case_char_code [code]:
+  "case_char f c = (let n = nat_of_char c in f (nibble_of_nat (n div 16)) (nibble_of_nat n))"
   by (cases c)
     (simp only: Let_def nibble_of_nat_of_char_div_16 nibble_of_nat_nat_of_char char.cases)
 
 lemma [code]:
-  "char_rec = char_case"
+  "rec_char = case_char"
   by (simp add: fun_eq_iff split: char.split)
 
 definition char_of_nat :: "nat \<Rightarrow> char" where
--- a/src/HOL/Sum_Type.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Sum_Type.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -85,6 +85,12 @@
   with assms show P by (auto simp add: sum_def Inl_def Inr_def)
 qed
 
+wrap_free_constructors [Inl, Inr] case_sum [isl] [[projl], [projr]]
+by (erule sumE, assumption) (auto dest: Inl_inject Inr_inject simp add: Inl_not_Inr)
+
+-- {* Avoid name clashes by prefixing the output of @{text rep_datatype} with @{text old}. *}
+setup {* Sign.mandatory_path "old" *}
+
 rep_datatype Inl Inr
 proof -
   fix P
@@ -93,6 +99,23 @@
   then show "P s" by (auto intro: sumE [of s])
 qed (auto dest: Inl_inject Inr_inject simp add: Inl_not_Inr)
 
+setup {* Sign.parent_path *}
+
+-- {* But erase the prefix for properties that are not generated by @{text wrap_free_constructors}. *}
+setup {* Sign.mandatory_path "sum" *}
+
+declare
+  old.sum.inject[iff del]
+  old.sum.distinct(1)[simp del, induct_simp del]
+
+lemmas induct = old.sum.induct
+lemmas inducts = old.sum.inducts
+lemmas recs = old.sum.recs
+lemmas cases = sum.case
+lemmas simps = sum.inject sum.distinct sum.case old.sum.recs
+
+setup {* Sign.parent_path *}
+
 primrec sum_map :: "('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> 'a + 'b \<Rightarrow> 'c + 'd" where
   "sum_map f1 f2 (Inl a) = Inl (f1 a)"
 | "sum_map f1 f2 (Inr a) = Inr (f2 a)"
@@ -123,44 +146,33 @@
 
 subsection {* Projections *}
 
-lemma sum_case_KK [simp]: "sum_case (\<lambda>x. a) (\<lambda>x. a) = (\<lambda>x. a)"
+lemma case_sum_KK [simp]: "case_sum (\<lambda>x. a) (\<lambda>x. a) = (\<lambda>x. a)"
   by (rule ext) (simp split: sum.split)
 
-lemma surjective_sum: "sum_case (\<lambda>x::'a. f (Inl x)) (\<lambda>y::'b. f (Inr y)) = f"
+lemma surjective_sum: "case_sum (\<lambda>x::'a. f (Inl x)) (\<lambda>y::'b. f (Inr y)) = f"
 proof
   fix s :: "'a + 'b"
   show "(case s of Inl (x\<Colon>'a) \<Rightarrow> f (Inl x) | Inr (y\<Colon>'b) \<Rightarrow> f (Inr y)) = f s"
     by (cases s) simp_all
 qed
 
-lemma sum_case_inject:
-  assumes a: "sum_case f1 f2 = sum_case g1 g2"
+lemma case_sum_inject:
+  assumes a: "case_sum f1 f2 = case_sum g1 g2"
   assumes r: "f1 = g1 \<Longrightarrow> f2 = g2 \<Longrightarrow> P"
   shows P
 proof (rule r)
   show "f1 = g1" proof
     fix x :: 'a
-    from a have "sum_case f1 f2 (Inl x) = sum_case g1 g2 (Inl x)" by simp
+    from a have "case_sum f1 f2 (Inl x) = case_sum g1 g2 (Inl x)" by simp
     then show "f1 x = g1 x" by simp
   qed
   show "f2 = g2" proof
     fix y :: 'b
-    from a have "sum_case f1 f2 (Inr y) = sum_case g1 g2 (Inr y)" by simp
+    from a have "case_sum f1 f2 (Inr y) = case_sum g1 g2 (Inr y)" by simp
     then show "f2 y = g2 y" by simp
   qed
 qed
 
-lemma sum_case_weak_cong:
-  "s = t \<Longrightarrow> sum_case f g s = sum_case f g t"
-  -- {* Prevents simplification of @{text f} and @{text g}: much faster. *}
-  by simp
-
-primrec Projl :: "'a + 'b \<Rightarrow> 'a" where
-  Projl_Inl: "Projl (Inl x) = x"
-
-primrec Projr :: "'a + 'b \<Rightarrow> 'b" where
-  Projr_Inr: "Projr (Inr x) = x"
-
 primrec Suml :: "('a \<Rightarrow> 'c) \<Rightarrow> 'a + 'b \<Rightarrow> 'c" where
   "Suml f (Inl x) = f x"
 
@@ -224,9 +236,6 @@
   } then show ?thesis by auto
 qed
 
-hide_const (open) Suml Sumr Projl Projr
-
-hide_const (open) sum
+hide_const (open) Suml Sumr sum
 
 end
-
--- a/src/HOL/Tools/BNF/bnf_def.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_def.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -12,7 +12,6 @@
   type nonemptiness_witness = {I: int list, wit: term, prop: thm list}
 
   val morph_bnf: morphism -> bnf -> bnf
-  val eq_bnf: bnf * bnf -> bool
   val bnf_of: Proof.context -> string -> bnf option
   val register_bnf: string -> (bnf * local_theory) -> (bnf * local_theory)
 
@@ -452,16 +451,12 @@
     wits = List.map (morph_witness phi) wits,
     rel = Morphism.term phi rel};
 
-fun eq_bnf (BNF {T = T1, live = live1, dead = dead1, ...},
-  BNF {T = T2, live = live2, dead = dead2, ...}) =
-  Type.could_unify (T1, T2) andalso live1 = live2 andalso dead1 = dead2;
-
 structure Data = Generic_Data
 (
   type T = bnf Symtab.table;
   val empty = Symtab.empty;
   val extend = I;
-  val merge = Symtab.merge eq_bnf;
+  fun merge data : T = Symtab.merge (K true) data;
 );
 
 fun bnf_of ctxt =
--- a/src/HOL/Tools/BNF/bnf_def_tactics.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_def_tactics.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -265,7 +265,7 @@
           rtac (thm RS ord_eq_le_trans) THEN' etac @{thm subset_trans[OF image_mono Un_upper1]})
         set_maps,
         rtac sym,
-        rtac (box_equals OF [map_cong0 OF replicate live @{thm fun_cong[OF sum_case_o_inj(1)]},
+        rtac (box_equals OF [map_cong0 OF replicate live @{thm fun_cong[OF case_sum_o_inj(1)]},
            map_comp RS sym, map_id])] 1
   end;
 
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -20,12 +20,12 @@
      co_iterss: term list list,
      mapss: thm list list,
      co_inducts: thm list,
+     co_inductss: thm list list,
      co_iter_thmsss: thm list list list,
      disc_co_itersss: thm list list list,
      sel_co_iterssss: thm list list list list};
 
   val of_fp_sugar: (fp_sugar -> 'a list) -> fp_sugar -> 'a
-  val eq_fp_sugar: fp_sugar * fp_sugar -> bool
   val morph_fp_sugar: morphism -> fp_sugar -> fp_sugar
   val transfer_fp_sugar: Proof.context -> fp_sugar -> fp_sugar
   val fp_sugar_of: Proof.context -> string -> fp_sugar option
@@ -95,7 +95,7 @@
   val co_datatypes: BNF_FP_Util.fp_kind -> (mixfix list -> binding list -> binding list ->
       binding list list -> binding list -> (string * sort) list -> typ list * typ list list ->
       BNF_Def.bnf list -> local_theory -> BNF_FP_Util.fp_result * local_theory) ->
-    (bool * (bool * bool)) * (((((binding * (typ * sort)) list * binding) * (binding * binding))
+    (bool * bool) * (((((binding * (typ * sort)) list * binding) * (binding * binding))
       * mixfix) * ((((binding * binding) * (binding * typ) list) * (binding * term) list) *
         mixfix) list) list ->
     local_theory -> local_theory
@@ -130,19 +130,16 @@
    co_iterss: term list list,
    mapss: thm list list,
    co_inducts: thm list,
+   co_inductss: thm list list,
    co_iter_thmsss: thm list list list,
    disc_co_itersss: thm list list list,
    sel_co_iterssss: thm list list list list};
 
 fun of_fp_sugar f (fp_sugar as ({index, ...}: fp_sugar)) = nth (f fp_sugar) index;
 
-fun eq_fp_sugar ({T = T1, fp = fp1, index = index1, fp_res = fp_res1, ...} : fp_sugar,
-    {T = T2, fp = fp2, index = index2, fp_res = fp_res2, ...} : fp_sugar) =
-  T1 = T2 andalso fp1 = fp2 andalso index1 = index2 andalso eq_fp_result (fp_res1, fp_res2);
-
 fun morph_fp_sugar phi ({T, fp, index, pre_bnfs, nested_bnfs, nesting_bnfs, fp_res, ctr_defss,
-    ctr_sugars, co_iterss, mapss, co_inducts, co_iter_thmsss, disc_co_itersss, sel_co_iterssss}
-    : fp_sugar) =
+    ctr_sugars, co_iterss, mapss, co_inducts, co_inductss, co_iter_thmsss, disc_co_itersss,
+    sel_co_iterssss} : fp_sugar) =
   {T = Morphism.typ phi T, fp = fp, index = index, pre_bnfs = map (morph_bnf phi) pre_bnfs,
     nested_bnfs = map (morph_bnf phi) nested_bnfs, nesting_bnfs = map (morph_bnf phi) nesting_bnfs,
    fp_res = morph_fp_result phi fp_res,
@@ -151,6 +148,7 @@
    co_iterss = map (map (Morphism.term phi)) co_iterss,
    mapss = map (map (Morphism.thm phi)) mapss,
    co_inducts = map (Morphism.thm phi) co_inducts,
+   co_inductss = map (map (Morphism.thm phi)) co_inductss,
    co_iter_thmsss = map (map (map (Morphism.thm phi))) co_iter_thmsss,
    disc_co_itersss = map (map (map (Morphism.thm phi))) disc_co_itersss,
    sel_co_iterssss = map (map (map (map (Morphism.thm phi)))) sel_co_iterssss};
@@ -163,7 +161,7 @@
   type T = fp_sugar Symtab.table;
   val empty = Symtab.empty;
   val extend = I;
-  val merge = Symtab.merge eq_fp_sugar;
+  fun merge data : T = Symtab.merge (K true) data;
 );
 
 fun fp_sugar_of ctxt =
@@ -183,14 +181,15 @@
     (fn phi => Data.map (Symtab.default (key, morph_fp_sugar phi fp_sugar)));
 
 fun register_fp_sugars fp pre_bnfs nested_bnfs nesting_bnfs (fp_res as {Ts, ...}) ctr_defss
-    ctr_sugars co_iterss mapss co_inducts co_iter_thmsss disc_co_itersss sel_co_iterssss lthy =
+    ctr_sugars co_iterss mapss co_inducts co_inductss co_iter_thmsss disc_co_itersss
+    sel_co_iterssss lthy =
   (0, lthy)
   |> fold (fn T as Type (s, _) => fn (kk, lthy) => (kk + 1,
     register_fp_sugar s {T = T, fp = fp, index = kk, pre_bnfs = pre_bnfs, nested_bnfs = nested_bnfs,
         nesting_bnfs = nesting_bnfs, fp_res = fp_res, ctr_defss = ctr_defss,
         ctr_sugars = ctr_sugars, co_iterss = co_iterss, mapss = mapss, co_inducts = co_inducts,
-        co_iter_thmsss = co_iter_thmsss, disc_co_itersss = disc_co_itersss,
-        sel_co_iterssss = sel_co_iterssss}
+        co_inductss = co_inductss, co_iter_thmsss = co_iter_thmsss,
+        disc_co_itersss = disc_co_itersss, sel_co_iterssss = sel_co_iterssss}
       lthy)) Ts
   |> snd;
 
@@ -211,6 +210,7 @@
 val id_def = @{thm id_def};
 val mp_conj = @{thm mp_conj};
 
+val fundefcong_attrs = @{attributes [fundef_cong]};
 val nitpicksimp_attrs = @{attributes [nitpick_simp]};
 val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
 val simp_attrs = @{attributes [simp]};
@@ -225,9 +225,7 @@
 fun flat_rec_arg_args xss =
   (* FIXME (once the old datatype package is phased out): The first line below gives the preferred
      order. The second line is for compatibility with the old datatype package. *)
-(*
-  flat xss
-*)
+  (* flat xss *)
   map hd xss @ maps tl xss;
 
 fun flat_corec_predss_getterss qss fss = maps (op @) (qss ~~ fss);
@@ -543,7 +541,7 @@
     fun generate_iter pre (_, _, fss, xssss) ctor_iter =
       (mk_binding pre,
        fold_rev (fold_rev Term.lambda) fss (Term.list_comb (ctor_iter,
-         map2 (mk_sum_caseN_balanced oo map2 mk_uncurried2_fun) fss xssss)));
+         map2 (mk_case_sumN_balanced oo map2 mk_uncurried2_fun) fss xssss)));
   in
     define_co_iters Least_FP fpT Cs (map3 generate_iter iterNs iter_args_typess' ctor_iters) lthy
   end;
@@ -670,7 +668,10 @@
             mk_induct_tac ctxt nn ns mss kksss (flat ctr_defss) ctor_induct' nested_set_maps
               pre_set_defss)
           |> singleton (Proof_Context.export names_lthy lthy)
-          |> Thm.close_derivation;
+          (* for "datatype_realizer.ML": *)
+          |> Thm.name_derivation (fst (dest_Type (hd fpTs)) ^ Long_Name.separator ^
+            (if nn > 1 then space_implode "_" (tl fp_b_names) ^ Long_Name.separator else "") ^
+            inductN);
       in
         `(conj_dests nn) thm
       end;
@@ -866,7 +867,7 @@
 
         fun tack z_name (c, u) f =
           let val z = Free (z_name, mk_sumT (fastype_of u, fastype_of c)) in
-            Term.lambda z (mk_sum_case (Term.lambda u u, Term.lambda c (f $ c)) $ z)
+            Term.lambda z (mk_case_sum (Term.lambda u u, Term.lambda c (f $ c)) $ z)
           end;
 
         fun build_coiter fcoiters maybe_mk_sumT maybe_tack cqf =
@@ -901,7 +902,7 @@
         val unfold_thmss = map2 (map2 prove) unfold_goalss unfold_tacss;
         val corec_thmss =
           map2 (map2 prove) corec_goalss corec_tacss
-          |> map (map (unfold_thms lthy @{thms sum_case_if}));
+          |> map (map (unfold_thms lthy @{thms case_sum_if}));
       in
         (unfold_thmss, corec_thmss)
       end;
@@ -977,7 +978,7 @@
   end;
 
 fun define_co_datatypes prepare_constraint prepare_typ prepare_term fp construct_fp
-    (wrap_opts as (no_discs_sels, (_, rep_compat)), specs) no_defs_lthy0 =
+    (wrap_opts as (no_discs_sels, _), specs) no_defs_lthy0 =
   let
     (* TODO: sanity checks on arguments *)
 
@@ -986,9 +987,6 @@
       else
         ();
 
-    fun qualify mandatory fp_b_name =
-      Binding.qualify mandatory fp_b_name o (rep_compat ? Binding.qualify false rep_compat_prefix);
-
     val nn = length specs;
     val fp_bs = map type_binding_of specs;
     val fp_b_names = map Binding.name_of fp_bs;
@@ -1040,7 +1038,7 @@
 
     val disc_bindingss = map (map disc_of) ctr_specss;
     val ctr_bindingss =
-      map2 (fn fp_b_name => map (qualify false fp_b_name o ctr_of)) fp_b_names ctr_specss;
+      map2 (fn fp_b_name => map (Binding.qualify false fp_b_name o ctr_of)) fp_b_names ctr_specss;
     val ctr_argsss = map (map args_of) ctr_specss;
     val ctr_mixfixess = map (map ctr_mixfix_of) ctr_specss;
 
@@ -1149,12 +1147,12 @@
     fun massage_simple_notes base =
       filter_out (null o #2)
       #> map (fn (thmN, thms, attrs) =>
-        ((qualify true base (Binding.name thmN), attrs), [(thms, [])]));
+        ((Binding.qualify true base (Binding.name thmN), attrs), [(thms, [])]));
 
     val massage_multi_notes =
       maps (fn (thmN, thmss, attrs) =>
         map3 (fn fp_b_name => fn Type (T_name, _) => fn thms =>
-            ((qualify true fp_b_name (Binding.name thmN), attrs T_name), [(thms, [])]))
+            ((Binding.qualify true fp_b_name (Binding.name thmN), attrs T_name), [(thms, [])]))
           fp_b_names fpTs thmss)
       #> filter_out (null o fst o hd o snd);
 
@@ -1224,8 +1222,8 @@
                     Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
                       mk_ctor_iff_dtor_tac ctxt (map (SOME o certifyT lthy) [dtorT, fpT])
                         (certify lthy ctor) (certify lthy dtor) ctor_dtor dtor_ctor)
+                    |> Morphism.thm phi
                     |> Thm.close_derivation
-                    |> Morphism.thm phi
                   end;
 
                 val sumEN_thm' =
@@ -1253,7 +1251,7 @@
               (sel_bindingss, sel_defaultss))) lthy
           end;
 
-        fun derive_maps_sets_rels (ctr_sugar, lthy) =
+        fun derive_maps_sets_rels (ctr_sugar as {case_cong, ...} : ctr_sugar, lthy) =
           if live = 0 then
             ((([], [], [], []), ctr_sugar), lthy)
           else
@@ -1327,7 +1325,8 @@
                 join_halves n half_rel_distinct_thmss other_half_rel_distinct_thmss;
 
               val anonymous_notes =
-                [(map (fn th => th RS @{thm eq_False[THEN iffD2]}) rel_distinct_thms,
+                [([case_cong], fundefcong_attrs),
+                 (map (fn th => th RS @{thm eq_False[THEN iffD2]}) rel_distinct_thms,
                   code_nitpicksimp_attrs),
                  (map2 (fn th => fn 0 => th RS @{thm eq_True[THEN iffD2]} | _ => th)
                     rel_inject_thms ms, code_nitpicksimp_attrs)]
@@ -1344,7 +1343,7 @@
                lthy |> Local_Theory.notes (anonymous_notes @ notes) |> snd)
             end;
 
-        fun mk_binding pre = qualify false fp_b_name (Binding.prefix_name (pre ^ "_") fp_b);
+        fun mk_binding pre = Binding.qualify false fp_b_name (Binding.prefix_name (pre ^ "_") fp_b);
 
         fun massage_res (((maps_sets_rels, ctr_sugar), co_iter_res), lthy) =
           (((maps_sets_rels, (ctrs, xss, ctr_defs, ctr_sugar)), co_iter_res), lthy);
@@ -1396,7 +1395,8 @@
         lthy
         |> Local_Theory.notes (common_notes @ notes) |> snd
         |> register_fp_sugars Least_FP pre_bnfs nested_bnfs nesting_bnfs fp_res ctr_defss ctr_sugars
-          iterss mapss [induct_thm] (transpose [fold_thmss, rec_thmss]) [] []
+          iterss mapss [induct_thm] (transpose [induct_thms]) (transpose [fold_thmss, rec_thmss])
+          [] []
       end;
 
     fun derive_note_coinduct_coiters_thms_for_types
@@ -1454,6 +1454,7 @@
         |> Local_Theory.notes (common_notes @ notes) |> snd
         |> register_fp_sugars Greatest_FP pre_bnfs nested_bnfs nesting_bnfs fp_res ctr_defss
           ctr_sugars coiterss mapss [coinduct_thm, strong_coinduct_thm]
+          (transpose [coinduct_thms, strong_coinduct_thms])
           (transpose [unfold_thmss, corec_thmss]) (transpose [disc_unfold_thmss, disc_corec_thmss])
           (transpose [sel_unfold_thmsss, sel_corec_thmsss])
       end;
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -84,8 +84,8 @@
   unfold_thms_tac ctxt @{thms sum.inject Pair_eq conj_assoc} THEN HEADGOAL (rtac refl);
 
 val iter_unfold_thms =
-  @{thms comp_def convol_def fst_conv id_def prod_case_Pair_iden snd_conv
-      split_conv unit_case_Unity} @ sum_prod_thms_map;
+  @{thms comp_def convol_def fst_conv id_def case_prod_Pair_iden snd_conv split_conv
+      case_unit_Unity} @ sum_prod_thms_map;
 
 fun mk_iter_tac pre_map_defs map_idents iter_defs ctor_iter ctr_def ctxt =
   unfold_thms_tac ctxt (ctr_def :: ctor_iter :: iter_defs @ pre_map_defs @ map_idents @
--- a/src/HOL/Tools/BNF/bnf_fp_n2m.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -69,7 +69,7 @@
     val dest_co_algT = co_swap o dest_funT;
     val co_alg_argT = fp_case fp range_type domain_type;
     val co_alg_funT = fp_case fp domain_type range_type;
-    val mk_co_product = curry (fp_case fp mk_convol mk_sum_case);
+    val mk_co_product = curry (fp_case fp mk_convol mk_case_sum);
     val mk_map_co_product = fp_case fp mk_prod_map mk_sum_map;
     val co_proj1_const = fp_case fp (fst_const o fst) (uncurry Inl_const o dest_sumT o snd);
     val mk_co_productT = curry (fp_case fp HOLogic.mk_prodT mk_sumT);
@@ -97,12 +97,12 @@
     val pre_bnfss = map #pre_bnfs fp_sugars;
     val nesty_bnfss = map (fn sugar => #nested_bnfs sugar @ #nesting_bnfs sugar) fp_sugars;
     val fp_nesty_bnfss = fp_bnfs :: nesty_bnfss;
-    val fp_nesty_bnfs = distinct eq_bnf (flat fp_nesty_bnfss);
+    val fp_nesty_bnfs = distinct (op = o pairself T_of_bnf) (flat fp_nesty_bnfss);
 
     val rels =
       let
         fun find_rel T As Bs = fp_nesty_bnfss
-          |> map (filter_out (curry eq_bnf BNF_Comp.DEADID_bnf))
+          |> map (filter_out (curry (op = o pairself name_of_bnf) BNF_Comp.DEADID_bnf))
           |> get_first (find_first (fn bnf => Type.could_unify (T_of_bnf bnf, T)))
           |> Option.map (fn bnf =>
             let val live = live_of_bnf bnf;
@@ -336,7 +336,7 @@
           o_apply comp_id id_comp map_pair.comp map_pair.id sum_map.comp sum_map.id};
         val rec_thms = fold_thms @ fp_case fp
           @{thms fst_convol map_pair_o_convol convol_o}
-          @{thms sum_case_o_inj(1) sum_case_o_sum_map o_sum_case};
+          @{thms case_sum_o_inj(1) case_sum_o_sum_map o_case_sum};
         val map_thms = no_refl (maps (fn bnf =>
           [map_comp0_of_bnf bnf RS sym, map_id0_of_bnf bnf]) fp_nesty_bnfs);
 
--- a/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -43,7 +43,7 @@
   type T = n2m_sugar Typtab.table;
   val empty = Typtab.empty;
   val extend = I;
-  fun merge data : T = Typtab.merge (eq_fst (eq_list eq_fp_sugar)) data;
+  fun merge data : T = Typtab.merge (K true) data;
 );
 
 fun morph_n2m_sugar phi (fp_sugars, (lfp_sugar_thms_opt, gfp_sugar_thms_opt)) =
@@ -64,7 +64,7 @@
 
 fun unfold_lets_splits (Const (@{const_name Let}, _) $ arg1 $ arg2) =
     unfold_lets_splits (betapply (arg2, arg1))
-  | unfold_lets_splits (t as Const (@{const_name prod_case}, _) $ u) =
+  | unfold_lets_splits (t as Const (@{const_name case_prod}, _) $ u) =
     (case unfold_lets_splits u of
       u' as Abs (s1, T1, Abs (s2, T2, _)) =>
       let val v = Var ((s1 ^ s2, Term.maxidx_of_term u' + 1), HOLogic.mk_prodT (T1, T2)) in
@@ -212,14 +212,14 @@
               (mk_binding b) fpTs Cs) fp_bs xtor_co_iterss lthy
           |>> split_list;
 
-        val ((co_inducts, un_fold_thmss, co_rec_thmss, disc_unfold_thmss, disc_corec_thmss,
-              sel_unfold_thmsss, sel_corec_thmsss), fp_sugar_thms) =
+        val ((co_inducts, co_inductss, un_fold_thmss, co_rec_thmss, disc_unfold_thmss,
+              disc_corec_thmss, sel_unfold_thmsss, sel_corec_thmsss), fp_sugar_thms) =
           if fp = Least_FP then
             derive_induct_iters_thms_for_types pre_bnfs (the iters_args_types) xtor_co_induct
               xtor_co_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss ctrss ctr_defss
               co_iterss co_iter_defss lthy
-            |> `(fn ((_, induct, _), (fold_thmss, rec_thmss, _)) =>
-              ([induct], fold_thmss, rec_thmss, [], [], [], []))
+            |> `(fn ((inducts, induct, _), (fold_thmss, rec_thmss, _)) =>
+              ([induct], [inducts], fold_thmss, rec_thmss, [], [], [], []))
             ||> (fn info => (SOME info, NONE))
           else
             derive_coinduct_coiters_thms_for_types pre_bnfs (the coiters_args_types) xtor_co_induct
@@ -229,8 +229,8 @@
             |> `(fn ((coinduct_thms_pairs, _), (unfold_thmss, corec_thmss, _),
                     (disc_unfold_thmss, disc_corec_thmss, _), _,
                     (sel_unfold_thmsss, sel_corec_thmsss, _)) =>
-              (map snd coinduct_thms_pairs, unfold_thmss, corec_thmss, disc_unfold_thmss,
-               disc_corec_thmss, sel_unfold_thmsss, sel_corec_thmsss))
+              (map snd coinduct_thms_pairs, map fst coinduct_thms_pairs, unfold_thmss, corec_thmss,
+               disc_unfold_thmss, disc_corec_thmss, sel_unfold_thmsss, sel_corec_thmsss))
             ||> (fn info => (NONE, SOME info));
 
         val phi = Proof_Context.export_morphism no_defs_lthy no_defs_lthy0;
@@ -239,6 +239,7 @@
           {T = T, fp = fp, index = kk, pre_bnfs = pre_bnfs, nested_bnfs = nested_bnfs,
            nesting_bnfs = nesting_bnfs, fp_res = fp_res, ctr_defss = ctr_defss,
            ctr_sugars = ctr_sugars, co_iterss = co_iterss, mapss = mapss, co_inducts = co_inducts,
+           co_inductss = transpose co_inductss,
            co_iter_thmsss = transpose [un_fold_thmss, co_rec_thmss],
            disc_co_itersss = transpose [disc_unfold_thmss, disc_corec_thmss],
            sel_co_iterssss = transpose [sel_unfold_thmsss, sel_corec_thmsss]}
--- a/src/HOL/Tools/BNF/bnf_fp_util.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_util.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -30,7 +30,6 @@
      rel_xtor_co_induct_thm: thm}
 
   val morph_fp_result: morphism -> fp_result -> fp_result
-  val eq_fp_result: fp_result * fp_result -> bool
   val un_fold_of: 'a list -> 'a
   val co_rec_of: 'a list -> 'a
 
@@ -145,13 +144,13 @@
   val Inl_const: typ -> typ -> term
   val Inr_const: typ -> typ -> term
 
+  val mk_case_sum: term * term -> term
+  val mk_case_sumN: term list -> term
+  val mk_case_sumN_balanced: term list -> term
   val mk_Inl: typ -> term -> term
   val mk_Inr: typ -> term -> term
   val mk_InN: typ list -> term -> int -> term
   val mk_InN_balanced: typ -> int -> term -> int -> term
-  val mk_sum_case: term * term -> term
-  val mk_sum_caseN: term list -> term
-  val mk_sum_caseN_balanced: term list -> term
 
   val dest_sumT: typ -> typ * typ
   val dest_sumTN: int -> typ -> typ list
@@ -167,8 +166,8 @@
   val mk_sumEN: int -> thm
   val mk_sumEN_balanced: int -> thm
   val mk_sumEN_tupled_balanced: int list -> thm
-  val mk_sum_casesN: int -> int -> thm
-  val mk_sum_casesN_balanced: int -> int -> thm
+  val mk_sum_caseN: int -> int -> thm
+  val mk_sum_caseN_balanced: int -> int -> thm
 
   val fixpoint: ('a * 'a -> bool) -> ('a list -> 'a list) -> 'a list -> 'a list
 
@@ -239,9 +238,6 @@
    xtor_co_iter_o_map_thmss = map (map (Morphism.thm phi)) xtor_co_iter_o_map_thmss,
    rel_xtor_co_induct_thm = Morphism.thm phi rel_xtor_co_induct_thm};
 
-fun eq_fp_result ({bnfs = bnfs1, ...} : fp_result, {bnfs = bnfs2, ...} : fp_result) =
-  eq_list eq_bnf (bnfs1, bnfs2);
-
 fun un_fold_of [f, _] = f;
 fun co_rec_of [_, r] = r;
 
@@ -413,17 +409,17 @@
     |> repair_types sum_T
   end;
 
-fun mk_sum_case (f, g) =
+fun mk_case_sum (f, g) =
   let
     val fT = fastype_of f;
     val gT = fastype_of g;
   in
-    Const (@{const_name sum_case},
+    Const (@{const_name case_sum},
       fT --> gT --> mk_sumT (domain_type fT, domain_type gT) --> range_type fT) $ f $ g
   end;
 
-val mk_sum_caseN = Library.foldr1 mk_sum_case;
-val mk_sum_caseN_balanced = Balanced_Tree.make mk_sum_case;
+val mk_case_sumN = Library.foldr1 mk_case_sum;
+val mk_case_sumN_balanced = Balanced_Tree.make mk_case_sum;
 
 fun If_const T = Const (@{const_name If}, HOLogic.boolT --> T --> T --> T);
 fun mk_If p t f = let val T = fastype_of t in If_const T $ p $ t $ f end;
@@ -474,18 +470,18 @@
     else mk_sumEN_balanced' n (map mk_tupled_allIN ms)
   end;
 
-fun mk_sum_casesN 1 1 = refl
-  | mk_sum_casesN _ 1 = @{thm sum.cases(1)}
-  | mk_sum_casesN 2 2 = @{thm sum.cases(2)}
-  | mk_sum_casesN n k = trans OF [@{thm sum_case_step(2)}, mk_sum_casesN (n - 1) (k - 1)];
+fun mk_sum_caseN 1 1 = refl
+  | mk_sum_caseN _ 1 = @{thm sum.case(1)}
+  | mk_sum_caseN 2 2 = @{thm sum.case(2)}
+  | mk_sum_caseN n k = trans OF [@{thm case_sum_step(2)}, mk_sum_caseN (n - 1) (k - 1)];
 
 fun mk_sum_step base step thm =
   if Thm.eq_thm_prop (thm, refl) then base else trans OF [step, thm];
 
-fun mk_sum_casesN_balanced 1 1 = refl
-  | mk_sum_casesN_balanced n k =
-    Balanced_Tree.access {left = mk_sum_step @{thm sum.cases(1)} @{thm sum_case_step(1)},
-      right = mk_sum_step @{thm sum.cases(2)} @{thm sum_case_step(2)}, init = refl} n k;
+fun mk_sum_caseN_balanced 1 1 = refl
+  | mk_sum_caseN_balanced n k =
+    Balanced_Tree.access {left = mk_sum_step @{thm sum.cases(1)} @{thm case_sum_step(1)},
+      right = mk_sum_step @{thm sum.cases(2)} @{thm case_sum_step(2)}, init = refl} n k;
 
 fun mk_rel_xtor_co_induct_thm fp pre_rels pre_phis rels phis xs ys xtors xtor's tac lthy =
   let
@@ -536,11 +532,11 @@
     val rewrite_comp_comp = fp_case fp @{thm rewriteR_comp_comp} @{thm rewriteL_comp_comp};
     val map_cong_passive_args1 = replicate m (fp_case fp @{thm id_comp} @{thm comp_id} RS fun_cong);
     val map_cong_active_args1 = replicate n (if is_rec
-      then fp_case fp @{thm convol_o} @{thm o_sum_case} RS fun_cong
+      then fp_case fp @{thm convol_o} @{thm o_case_sum} RS fun_cong
       else refl);
     val map_cong_passive_args2 = replicate m (fp_case fp @{thm comp_id} @{thm id_comp} RS fun_cong);
     val map_cong_active_args2 = replicate n (if is_rec
-      then fp_case fp @{thm map_pair_o_convol_id} @{thm sum_case_o_sum_map_id}
+      then fp_case fp @{thm map_pair_o_convol_id} @{thm case_sum_o_sum_map_id}
       else fp_case fp @{thm id_comp} @{thm comp_id} RS fun_cong);
     fun mk_map_congs passive active = map (fn thm => thm OF (passive @ active) RS ext) map_cong0s;
     val map_cong1s = mk_map_congs map_cong_passive_args1 map_cong_active_args1;
--- a/src/HOL/Tools/BNF/bnf_gfp.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -493,16 +493,16 @@
         |> Thm.close_derivation
       end;
 
-    val mor_sum_case_thm =
+    val mor_case_sum_thm =
       let
         val maps = map3 (fn s => fn sum_s => fn mapx =>
-          mk_sum_case (HOLogic.mk_comp (Term.list_comb (mapx, passive_ids @ Inls), s), sum_s))
+          mk_case_sum (HOLogic.mk_comp (Term.list_comb (mapx, passive_ids @ Inls), s), sum_s))
           s's sum_ss map_Inls;
       in
         Goal.prove_sorry lthy [] []
           (fold_rev Logic.all (s's @ sum_ss) (HOLogic.mk_Trueprop
             (mk_mor (map HOLogic.mk_UNIV activeBs) s's sum_UNIVs maps Inls)))
-          (K (mk_mor_sum_case_tac ks mor_UNIV_thm))
+          (K (mk_mor_case_sum_tac ks mor_UNIV_thm))
         |> Thm.close_derivation
       end;
 
@@ -525,7 +525,7 @@
         val Suc = Term.absdummy HOLogic.natT (Term.absfree hrec'
           (HOLogic.mk_tuple (map4 mk_Suc ss setssAs zs zs')));
 
-        val rhs = mk_nat_rec Zero Suc;
+        val rhs = mk_rec_nat Zero Suc;
       in
         fold_rev (Term.absfree o Term.dest_Free) ss rhs
       end;
@@ -555,8 +555,8 @@
         mk_nthN n (Term.list_comb (Const (nth hset_recs (j - 1), hset_recT), args)) i
       end;
 
-    val hset_rec_0ss = mk_rec_simps n @{thm nat_rec_0_imp} hset_rec_defs;
-    val hset_rec_Sucss = mk_rec_simps n @{thm nat_rec_Suc_imp} hset_rec_defs;
+    val hset_rec_0ss = mk_rec_simps n @{thm rec_nat_0_imp} hset_rec_defs;
+    val hset_rec_Sucss = mk_rec_simps n @{thm rec_nat_Suc_imp} hset_rec_defs;
     val hset_rec_0ss' = transpose hset_rec_0ss;
     val hset_rec_Sucss' = transpose hset_rec_Sucss;
 
@@ -1133,7 +1133,7 @@
         val fs = map mk_undefined fTs1 @ (f :: map mk_undefined fTs2);
       in
         HOLogic.mk_split (Term.absfree Kl' (Term.absfree lab'
-          (mk_sum_caseN fs $ (lab $ HOLogic.mk_list sum_sbdT []))))
+          (mk_case_sumN fs $ (lab $ HOLogic.mk_list sum_sbdT []))))
       end;
 
     val ((strT_frees, (_, strT_def_frees)), (lthy, lthy_old)) =
@@ -1202,7 +1202,7 @@
         val Suc = Term.absdummy HOLogic.natT (Term.absfree Lev_rec'
           (HOLogic.mk_tuple (map5 mk_Suc ks ss setssAs zs zs')));
 
-        val rhs = mk_nat_rec Zero Suc;
+        val rhs = mk_rec_nat Zero Suc;
       in
         fold_rev (Term.absfree o Term.dest_Free) ss rhs
       end;
@@ -1226,8 +1226,8 @@
         mk_nthN n (Term.list_comb (Const (Lev, LevT), ss) $ nat) i
       end;
 
-    val Lev_0s = flat (mk_rec_simps n @{thm nat_rec_0_imp} [Lev_def]);
-    val Lev_Sucs = flat (mk_rec_simps n @{thm nat_rec_Suc_imp} [Lev_def]);
+    val Lev_0s = flat (mk_rec_simps n @{thm rec_nat_0_imp} [Lev_def]);
+    val Lev_Sucs = flat (mk_rec_simps n @{thm rec_nat_Suc_imp} [Lev_def]);
 
     val rv_bind = mk_internal_b rvN;
     val rv_def_bind = rpair [] (Thm.def_binding rv_bind);
@@ -1239,13 +1239,13 @@
             fun mk_case i' =
               Term.absfree k' (mk_nthN n rv_rec i' $ (mk_from_sbd s b i i' $ k));
           in
-            Term.absfree b' (mk_sum_caseN (map mk_case ks) $ sumx)
+            Term.absfree b' (mk_case_sumN (map mk_case ks) $ sumx)
           end;
 
         val Cons = Term.absfree sumx' (Term.absdummy sum_sbd_listT (Term.absfree rv_rec'
           (HOLogic.mk_tuple (map4 mk_Cons ks ss zs zs'))));
 
-        val rhs = mk_list_rec Nil Cons;
+        val rhs = mk_rec_list Nil Cons;
       in
         fold_rev (Term.absfree o Term.dest_Free) ss rhs
       end;
@@ -1270,8 +1270,8 @@
         mk_nthN n (Term.list_comb (Const (rv, rvT), ss) $ kl) i
       end;
 
-    val rv_Nils = flat (mk_rec_simps n @{thm list_rec_Nil_imp} [rv_def]);
-    val rv_Conss = flat (mk_rec_simps n @{thm list_rec_Cons_imp} [rv_def]);
+    val rv_Nils = flat (mk_rec_simps n @{thm rec_list_Nil_imp} [rv_def]);
+    val rv_Conss = flat (mk_rec_simps n @{thm rec_list_Cons_imp} [rv_def]);
 
     val beh_binds = mk_internal_bs behN;
     fun beh_bind i = nth beh_binds (i - 1);
@@ -1285,7 +1285,7 @@
 
         val Lab = Term.absfree kl' (mk_If
           (HOLogic.mk_mem (kl, mk_Lev ss (mk_size kl) i $ z))
-          (mk_sum_caseN (map5 mk_case ks to_sbd_maps ss zs zs') $ (mk_rv ss kl i $ z))
+          (mk_case_sumN (map5 mk_case ks to_sbd_maps ss zs zs') $ (mk_rv ss kl i $ z))
           (mk_undefined sbdFT));
 
         val rhs = HOLogic.mk_prod (mk_UNION (HOLogic.mk_UNIV HOLogic.natT)
@@ -1413,7 +1413,7 @@
 
         fun mk_conjunct i z B = HOLogic.mk_imp
           (HOLogic.mk_conj (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), HOLogic.mk_mem (z, B)),
-          mk_sum_caseN (map4 mk_case ss setssAs zs zs') $ (mk_rv ss kl i $ z));
+          mk_case_sumN (map4 mk_case ss setssAs zs zs') $ (mk_rv ss kl i $ z));
 
         val goal = list_all_free (kl :: zs)
           (Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct ks zs Bs));
@@ -1432,8 +1432,8 @@
         map (fn i => map (fn i' =>
           split_conj_thm (if n = 1 then set_rv_Lev' RS mk_conjunctN n i RS mp
             else set_rv_Lev' RS mk_conjunctN n i RS mp RSN
-              (2, @{thm sum_case_weak_cong} RS iffD1) RS
-              (mk_sum_casesN n i' RS iffD1))) ks) ks
+              (2, @{thm sum.weak_case_cong} RS iffD1) RS
+              (mk_sum_caseN n i' RS iffD1))) ks) ks
       end;
 
     val set_Lev_thmsss =
@@ -1828,7 +1828,7 @@
 
     val corec_Inl_sum_thms =
       let
-        val mor = mor_comp_thm OF [mor_sum_case_thm, mor_unfold_thm];
+        val mor = mor_comp_thm OF [mor_case_sum_thm, mor_unfold_thm];
       in
         map2 (fn unique => fn unfold_dtor =>
           trans OF [mor RS unique, unfold_dtor]) unfold_unique_mor_thms unfold_dtor_thms
@@ -1839,7 +1839,7 @@
 
     val corec_strs =
       map3 (fn dtor => fn sum_s => fn mapx =>
-        mk_sum_case
+        mk_case_sum
           (HOLogic.mk_comp (Term.list_comb (mapx, passive_ids @ corec_Inls), dtor), sum_s))
       dtors corec_ss corec_maps;
 
@@ -1863,14 +1863,14 @@
     val corec_defs = map (fn def =>
       mk_unabs_def n (Morphism.thm phi def RS meta_eq_to_obj_eq)) corec_def_frees;
 
-    val sum_cases =
-      map2 (fn T => fn i => mk_sum_case (HOLogic.id_const T, mk_corec corec_ss i)) Ts ks;
+    val case_sums =
+      map2 (fn T => fn i => mk_case_sum (HOLogic.id_const T, mk_corec corec_ss i)) Ts ks;
     val dtor_corec_thms =
       let
         fun mk_goal i corec_s corec_map dtor z =
           let
             val lhs = dtor $ (mk_corec corec_ss i $ z);
-            val rhs = Term.list_comb (corec_map, passive_ids @ sum_cases) $ (corec_s $ z);
+            val rhs = Term.list_comb (corec_map, passive_ids @ case_sums) $ (corec_s $ z);
           in
             fold_rev Logic.all (z :: corec_ss) (mk_Trueprop_eq (lhs, rhs))
           end;
@@ -1886,7 +1886,7 @@
 
     val corec_unique_mor_thm =
       let
-        val id_fs = map2 (fn T => fn f => mk_sum_case (HOLogic.id_const T, f)) Ts unfold_fs;
+        val id_fs = map2 (fn T => fn f => mk_case_sum (HOLogic.id_const T, f)) Ts unfold_fs;
         val prem = HOLogic.mk_Trueprop (mk_mor corec_UNIVs corec_strs UNIVs dtors id_fs);
         fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_corec corec_ss i);
         val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
@@ -1907,9 +1907,9 @@
     val (dtor_corec_unique_thms, dtor_corec_unique_thm) =
       `split_conj_thm (split_conj_prems n
         (mor_UNIV_thm RS iffD2 RS corec_unique_mor_thm)
-        |> Local_Defs.unfold lthy (@{thms o_sum_case comp_id id_comp comp_assoc[symmetric]
-           sum_case_o_inj(1)} @ map_id0s_o_id @ sym_map_comps)
-        OF replicate n @{thm arg_cong2[of _ _ _ _ sum_case, OF refl]});
+        |> Local_Defs.unfold lthy (@{thms o_case_sum comp_id id_comp comp_assoc[symmetric]
+           case_sum_o_inj(1)} @ map_id0s_o_id @ sym_map_comps)
+        OF replicate n @{thm arg_cong2[of _ _ _ _ case_sum, OF refl]});
 
     val timer = time (timer "corec definitions & thms");
 
--- a/src/HOL/Tools/BNF/bnf_gfp_rec_sugar.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -195,7 +195,7 @@
           let val branches' = map (massage_rec bound_Ts) branches in
             Term.list_comb (If_const (typof (hd branches')) $ tap check_no_call obj, branches')
           end
-        | (c as Const (@{const_name prod_case}, _), arg :: args) =>
+        | (c as Const (@{const_name case_prod}, _), arg :: args) =>
           massage_rec bound_Ts
             (unfold_lets_splits (Term.list_comb (c $ Envir.eta_long bound_Ts arg, args)))
         | (Const (c, _), args as _ :: _ :: _) =>
@@ -295,12 +295,12 @@
               end
             | NONE =>
               (case t of
-                Const (@{const_name prod_case}, _) $ t' =>
+                Const (@{const_name case_prod}, _) $ t' =>
                 let
                   val U' = curried_type U;
                   val T' = curried_type T;
                 in
-                  Const (@{const_name prod_case}, U' --> U) $ massage_call bound_Ts U' T' t'
+                  Const (@{const_name case_prod}, U' --> U) $ massage_call bound_Ts U' T' t'
                 end
               | t1 $ t2 =>
                 (if has_call t2 then
@@ -340,22 +340,8 @@
 fun fold_rev_corec_code_rhs ctxt f =
   fold_rev_let_if_case ctxt (fn conds => uncurry (f conds) o Term.strip_comb);
 
-fun case_thms_of_term ctxt bound_Ts t =
-  let
-    fun ctr_sugar_of_case c s =
-      (case ctr_sugar_of ctxt s of
-        SOME (ctr_sugar as {casex = Const (c', _), sel_splits = _ :: _, ...}) =>
-        if c' = c then SOME ctr_sugar else NONE
-      | _ => NONE);
-    fun add_ctr_sugar (s, Type (@{type_name fun}, [_, T])) =
-        binder_types T
-        |> map_filter (try (fst o dest_Type))
-        |> distinct (op =)
-        |> map_filter (ctr_sugar_of_case s)
-      | add_ctr_sugar _ = [];
-
-    val ctr_sugars = maps add_ctr_sugar (Term.add_consts t []);
-  in
+fun case_thms_of_term ctxt t =
+  let val ctr_sugars = map_filter (Ctr_Sugar.ctr_sugar_of_case ctxt o fst) (Term.add_consts t []) in
     (maps #distincts ctr_sugars, maps #discIs ctr_sugars, maps #disc_exhausts ctr_sugars,
      maps #sel_splits ctr_sugars, maps #sel_split_asms ctr_sugars)
   end;
@@ -785,7 +771,7 @@
               let val (u, vs) = strip_comb t in
                 if is_Free u andalso has_call u then
                   Inr_const U T $ mk_tuple1 bound_Ts vs
-                else if try (fst o dest_Const) u = SOME @{const_name prod_case} then
+                else if try (fst o dest_Const) u = SOME @{const_name case_prod} then
                   map (rewrite bound_Ts) vs |> chop 1
                   |>> HOLogic.mk_split o the_single
                   |> Term.list_comb
@@ -1020,12 +1006,12 @@
         de_facto_exhaustives disc_eqnss
       |> list_all_fun_args []
     val nchotomy_taut_thmss =
-      map6 (fn tac_opt => fn {disc_exhausts = res_disc_exhausts, ...} => fn arg_Ts =>
+      map5 (fn tac_opt => fn {disc_exhausts = res_disc_exhausts, ...} =>
           fn {code_rhs_opt, ...} :: _ => fn [] => K []
             | [goal] => fn true =>
               let
                 val (_, _, arg_disc_exhausts, _, _) =
-                  case_thms_of_term lthy arg_Ts (the_default Term.dummy code_rhs_opt);
+                  case_thms_of_term lthy (the_default Term.dummy code_rhs_opt);
               in
                 [Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
                    mk_primcorec_nchotomy_tac ctxt (res_disc_exhausts @ arg_disc_exhausts))
@@ -1035,7 +1021,7 @@
               (case tac_opt of
                 SOME tac => [Goal.prove_sorry lthy [] [] goal tac |> Thm.close_derivation]
               | NONE => []))
-        tac_opts corec_specs arg_Tss disc_eqnss nchotomy_goalss syntactic_exhaustives;
+        tac_opts corec_specs disc_eqnss nchotomy_goalss syntactic_exhaustives;
 
     val syntactic_exhaustives =
       map (fn disc_eqns => forall (null o #prems orf is_some o #code_rhs_opt) disc_eqns
@@ -1132,7 +1118,7 @@
               |> HOLogic.mk_Trueprop o HOLogic.mk_eq
               |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
               |> curry Logic.list_all (map dest_Free fun_args);
-            val (distincts, _, _, sel_splits, sel_split_asms) = case_thms_of_term lthy [] rhs_term;
+            val (distincts, _, _, sel_splits, sel_split_asms) = case_thms_of_term lthy rhs_term;
           in
             mk_primcorec_sel_tac lthy def_thms distincts sel_splits sel_split_asms nested_maps
               nested_map_idents nested_map_comps sel_corec k m excludesss
@@ -1258,7 +1244,7 @@
                       |> pairself (curry mk_Trueprop_eq (applied_fun_of fun_name fun_T fun_args)
                         #> curry Logic.list_all (map dest_Free fun_args));
                     val (distincts, discIs, _, sel_splits, sel_split_asms) =
-                      case_thms_of_term lthy bound_Ts raw_rhs;
+                      case_thms_of_term lthy raw_rhs;
 
                     val raw_code_thm = mk_primcorec_raw_code_tac lthy distincts discIs sel_splits
                         sel_split_asms ms ctr_thms
--- a/src/HOL/Tools/BNF/bnf_gfp_tactics.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp_tactics.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -57,13 +57,13 @@
     thm list -> thm list -> thm list -> thm list -> thm list list -> thm list list list ->
     thm list list list -> thm list list list -> thm list list -> thm list list -> thm list ->
     thm list -> thm list -> tactic
+  val mk_mor_case_sum_tac: 'a list -> thm -> tactic
   val mk_mor_comp_tac: thm -> thm list -> thm list -> thm list -> tactic
   val mk_mor_elim_tac: thm -> tactic
   val mk_mor_hset_rec_tac: int -> int -> cterm option list -> int -> thm list -> thm list ->
     thm list -> thm list list -> thm list list -> tactic
   val mk_mor_incl_tac: thm -> thm list -> tactic
   val mk_mor_str_tac: 'a list -> thm -> tactic
-  val mk_mor_sum_case_tac: 'a list -> thm -> tactic
   val mk_mor_unfold_tac: int -> thm -> thm list -> thm list -> thm list -> thm list -> thm list ->
     thm list -> tactic
   val mk_prefCl_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list -> tactic
@@ -112,7 +112,7 @@
   @{thm ord_eq_le_trans[OF trans[OF fun_cong[OF image_id] id_apply]]};
 val ordIso_ordLeq_trans = @{thm ordIso_ordLeq_trans};
 val snd_convol_fun_cong_sym = @{thm snd_convol} RS fun_cong RS sym;
-val sum_case_weak_cong = @{thm sum_case_weak_cong};
+val sum_weak_case_cong = @{thm sum.weak_case_cong};
 val trans_fun_cong_image_id_id_apply = @{thm trans[OF fun_cong[OF image_id] id_apply]};
 val Collect_splitD_set_mp = @{thm Collect_splitD[OF set_mp]};
 val rev_bspec = Drule.rotate_prems 1 bspec;
@@ -170,8 +170,8 @@
 fun mk_mor_str_tac ks mor_UNIV =
   (stac mor_UNIV THEN' CONJ_WRAP' (K (rtac refl)) ks) 1;
 
-fun mk_mor_sum_case_tac ks mor_UNIV =
-  (stac mor_UNIV THEN' CONJ_WRAP' (K (rtac @{thm sum_case_o_inj(1)[symmetric]})) ks) 1;
+fun mk_mor_case_sum_tac ks mor_UNIV =
+  (stac mor_UNIV THEN' CONJ_WRAP' (K (rtac @{thm case_sum_o_inj(1)[symmetric]})) ks) 1;
 
 fun mk_set_incl_hset_tac def rec_Suc =
   EVERY' (stac def ::
@@ -376,7 +376,7 @@
     fun coalg_tac (i, ((passive_sets, active_sets), def)) =
       EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
         hyp_subst_tac ctxt, rtac (def RS trans RS @{thm ssubst_mem}), etac (arg_cong RS trans),
-        rtac (mk_sum_casesN n i), rtac CollectI,
+        rtac (mk_sum_caseN n i), rtac CollectI,
         EVERY' (map (fn thm => EVERY' [rtac conjI, rtac (thm RS ord_eq_le_trans),
           etac ((trans OF [@{thm image_id} RS fun_cong, id_apply]) RS ord_eq_le_trans)])
           passive_sets),
@@ -504,7 +504,7 @@
       CONJ_WRAP' (fn rv_Cons =>
         CONJ_WRAP' (fn (i, rv_Nil) => (EVERY' [rtac exI,
           rtac (@{thm append_Nil} RS arg_cong RS trans),
-          rtac (rv_Cons RS trans), rtac (mk_sum_casesN n i RS arg_cong RS trans), rtac rv_Nil]))
+          rtac (rv_Cons RS trans), rtac (mk_sum_caseN n i RS arg_cong RS trans), rtac rv_Nil]))
         (ks ~~ rv_Nils))
       rv_Conss,
       REPEAT_DETERM o rtac allI, rtac (mk_sumEN n),
@@ -512,8 +512,8 @@
         CONJ_WRAP' (fn rv_Cons => EVERY' [REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i),
           CONJ_WRAP' (fn i' => EVERY' [dtac (mk_conjunctN n i'), etac exE, rtac exI,
             rtac (@{thm append_Cons} RS arg_cong RS trans),
-            rtac (rv_Cons RS trans), etac (sum_case_weak_cong RS arg_cong RS trans),
-            rtac (mk_sum_casesN n i RS arg_cong RS trans), atac])
+            rtac (rv_Cons RS trans), etac (sum_weak_case_cong RS arg_cong RS trans),
+            rtac (mk_sum_caseN n i RS arg_cong RS trans), atac])
           ks])
         rv_Conss)
       ks)] 1
@@ -530,7 +530,7 @@
         EVERY' [rtac impI, REPEAT_DETERM o etac conjE,
           dtac (Lev_0 RS equalityD1 RS set_mp), etac @{thm singletonE}, etac ssubst,
           rtac (rv_Nil RS arg_cong RS iffD2),
-          rtac (mk_sum_casesN n i RS iffD2),
+          rtac (mk_sum_caseN n i RS iffD2),
           CONJ_WRAP' (fn thm => etac thm THEN' atac) (take m coalg_sets)])
       (ks ~~ ((Lev_0s ~~ rv_Nils) ~~ coalg_setss)),
       REPEAT_DETERM o rtac allI,
@@ -540,7 +540,7 @@
             (fn (i, (from_to_sbd, coalg_set)) =>
               EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt,
               rtac (rv_Cons RS arg_cong RS iffD2),
-              rtac (mk_sum_casesN n i RS arg_cong RS trans RS iffD2),
+              rtac (mk_sum_caseN n i RS arg_cong RS trans RS iffD2),
               etac (from_to_sbd RS arg_cong), REPEAT_DETERM o etac allE,
               dtac (mk_conjunctN n i), etac mp, etac conjI, etac set_rev_mp,
               etac coalg_set, atac])
@@ -583,7 +583,7 @@
                       rtac @{thm ssubst_mem[OF append_Cons]}, rtac (mk_UnIN n i),
                       rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl,
                       rtac conjI, atac, dtac (sym RS trans RS sym),
-                      rtac (rv_Cons RS trans), rtac (mk_sum_casesN n i RS trans),
+                      rtac (rv_Cons RS trans), rtac (mk_sum_caseN n i RS trans),
                       etac (from_to_sbd RS arg_cong), REPEAT_DETERM o etac allE,
                       dtac (mk_conjunctN n i), dtac mp, atac,
                       dtac (mk_conjunctN n i'), dtac mp, atac,
@@ -639,7 +639,7 @@
                     atac, atac, hyp_subst_tac ctxt] THEN'
                     CONJ_WRAP' (fn i'' =>
                       EVERY' [rtac impI, dtac (sym RS trans),
-                        rtac (rv_Cons RS trans), rtac (mk_sum_casesN n i RS arg_cong RS trans),
+                        rtac (rv_Cons RS trans), rtac (mk_sum_caseN n i RS arg_cong RS trans),
                         etac (from_to_sbd RS arg_cong),
                         REPEAT_DETERM o etac allE,
                         dtac (mk_conjunctN n i), dtac mp, atac,
@@ -684,7 +684,7 @@
               rtac exI, rtac conjI,
               (if n = 1 then rtac @{thm if_P} THEN' etac length_Lev'
               else rtac (@{thm if_P} RS arg_cong RS trans) THEN' etac length_Lev' THEN'
-                etac (sum_case_weak_cong RS trans) THEN' rtac (mk_sum_casesN n i)),
+                etac (sum_weak_case_cong RS trans) THEN' rtac (mk_sum_caseN n i)),
               EVERY' (map2 (fn set_map0 => fn set_rv_Lev =>
                 EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac (set_map0 RS trans),
                   rtac trans_fun_cong_image_id_id_apply,
@@ -708,7 +708,7 @@
               rtac exI, rtac conjI,
               (if n = 1 then rtac @{thm if_P} THEN' etac length_Lev'
               else rtac (@{thm if_P} RS trans) THEN' etac length_Lev' THEN'
-                etac (sum_case_weak_cong RS trans) THEN' rtac (mk_sum_casesN n i)),
+                etac (sum_weak_case_cong RS trans) THEN' rtac (mk_sum_caseN n i)),
               EVERY' (map2 (fn set_map0 => fn set_rv_Lev =>
                 EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac (set_map0 RS trans),
                   rtac trans_fun_cong_image_id_id_apply,
@@ -735,7 +735,7 @@
           rtac length_Lev', rtac (Lev_0 RS equalityD2 RS set_mp), rtac @{thm singletonI},
           CONVERSION (Conv.top_conv
             (K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt),
-          if n = 1 then rtac refl else rtac (mk_sum_casesN n i),
+          if n = 1 then rtac refl else rtac (mk_sum_caseN n i),
           EVERY' (map2 (fn set_map0 => fn coalg_set =>
             EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac (set_map0 RS trans),
               rtac trans_fun_cong_image_id_id_apply, etac coalg_set, atac])
@@ -756,15 +756,14 @@
     fun mor_tac (i, (strT_def, (((Lev_0, Lev_Suc), (rv_Nil, rv_Cons)),
       ((map_comp_id, (map_cong0, map_arg_cong)), (length_Lev', (from_to_sbds, to_sbd_injs)))))) =
       EVERY' [rtac ballI, rtac sym, rtac trans, rtac strT_def,
-        rtac (@{thm if_P} RS
-          (if n = 1 then map_arg_cong else sum_case_weak_cong) RS trans),
+        rtac (@{thm if_P} RS (if n = 1 then map_arg_cong else sum_weak_case_cong) RS trans),
         rtac (@{thm list.size(3)} RS arg_cong RS trans RS equalityD2 RS set_mp),
         rtac Lev_0, rtac @{thm singletonI},
         CONVERSION (Conv.top_conv
           (K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt),
         if n = 1 then K all_tac
-        else (rtac (sum_case_weak_cong RS trans) THEN'
-          rtac (mk_sum_casesN n i) THEN' rtac (mk_sum_casesN n i RS trans)),
+        else (rtac (sum_weak_case_cong RS trans) THEN'
+          rtac (mk_sum_caseN n i) THEN' rtac (mk_sum_caseN n i RS trans)),
         rtac (map_comp_id RS trans), rtac (map_cong0 OF replicate m refl),
         EVERY' (map3 (fn i' => fn to_sbd_inj => fn from_to_sbd =>
           DETERM o EVERY' [rtac trans, rtac o_apply, rtac Pair_eqI, rtac conjI,
@@ -801,7 +800,7 @@
             CONVERSION (Conv.top_conv
               (K (Conv.try_conv (Conv.rewr_conv (rv_Cons RS eq_reflection)))) ctxt),
             if n = 1 then K all_tac
-            else rtac sum_case_weak_cong THEN' rtac (mk_sum_casesN n i' RS trans),
+            else rtac sum_weak_case_cong THEN' rtac (mk_sum_caseN n i' RS trans),
             SELECT_GOAL (unfold_thms_tac ctxt [from_to_sbd]), rtac refl,
             rtac refl])
         ks to_sbd_injs from_to_sbds)];
@@ -928,11 +927,11 @@
   unfold_thms_tac ctxt corec_defs THEN EVERY' [rtac trans, rtac (o_apply RS arg_cong),
     rtac trans, rtac unfold, fo_rtac (@{thm sum.cases(2)} RS arg_cong RS trans) ctxt, rtac map_cong0,
     REPEAT_DETERM_N m o rtac refl,
-    EVERY' (map (fn thm => rtac @{thm sum_case_expand_Inr} THEN' rtac thm) corec_Inls)] 1;
+    EVERY' (map (fn thm => rtac @{thm case_sum_expand_Inr} THEN' rtac thm) corec_Inls)] 1;
 
 fun mk_corec_unique_mor_tac ctxt corec_defs corec_Inls unfold_unique_mor =
   unfold_thms_tac ctxt
-    (corec_defs @ map (fn thm => thm RS @{thm sum_case_expand_Inr'}) corec_Inls) THEN
+    (corec_defs @ map (fn thm => thm RS @{thm case_sum_expand_Inr'}) corec_Inls) THEN
   etac unfold_unique_mor 1;
 
 fun mk_dtor_coinduct_tac m raw_coind bis_rel rel_congs =
@@ -947,7 +946,7 @@
     rel_congs,
     rtac impI, REPEAT_DETERM o etac conjE,
     CONJ_WRAP' (K (EVERY' [rtac impI, rtac @{thm IdD}, etac set_mp,
-      rtac CollectI, etac @{thm prod_caseI}])) rel_congs] 1;
+      rtac CollectI, etac @{thm case_prodI}])) rel_congs] 1;
 
 fun mk_dtor_map_coinduct_tac m ks raw_coind bis_def =
   let
@@ -1141,7 +1140,7 @@
         passive_set_map0s dtor_set_incls),
         CONJ_WRAP' (fn (in_Jrel, (set_map0, dtor_set_set_incls)) =>
           EVERY' [rtac ord_eq_le_trans, rtac set_map0, rtac @{thm image_subsetI}, rtac CollectI,
-            rtac @{thm prod_caseI}, rtac (in_Jrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+            rtac @{thm case_prodI}, rtac (in_Jrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
             CONJ_WRAP' (fn thm => etac (thm RS @{thm subset_trans}) THEN' atac) dtor_set_set_incls,
             rtac conjI, rtac refl, rtac refl])
         (in_Jrels ~~ (active_set_map0s ~~ dtor_set_set_inclss)),
@@ -1165,7 +1164,7 @@
                 dtac set_rev_mp, etac @{thm image_mono}, etac imageE,
                 dtac @{thm ssubst_mem[OF pair_collapse]},
                 REPEAT_DETERM o eresolve_tac (CollectE :: conjE ::
-                  @{thms prod_caseE iffD1[OF Pair_eq, elim_format]}),
+                  @{thms case_prodE iffD1[OF Pair_eq, elim_format]}),
                 hyp_subst_tac ctxt,
                 dtac (in_Jrel RS iffD1),
                 dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE,
@@ -1181,7 +1180,7 @@
           EVERY' (map (fn in_Jrel => EVERY' [rtac trans, rtac o_apply, dtac set_rev_mp, atac,
             dtac @{thm ssubst_mem[OF pair_collapse]},
             REPEAT_DETERM o
-              eresolve_tac (CollectE :: conjE :: @{thms prod_caseE iffD1[OF Pair_eq, elim_format]}),
+              eresolve_tac (CollectE :: conjE :: @{thms case_prodE iffD1[OF Pair_eq, elim_format]}),
             hyp_subst_tac ctxt, dtac (in_Jrel RS iffD1),
             dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, atac]) in_Jrels),
           atac]]
--- a/src/HOL/Tools/BNF/bnf_gfp_util.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp_util.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -22,12 +22,12 @@
   val mk_in_rel: term -> term
   val mk_equiv: term -> term -> term
   val mk_fromCard: term -> term -> term
-  val mk_list_rec: term -> term -> term
-  val mk_nat_rec: term -> term -> term
   val mk_prefCl: term -> term
   val mk_prefixeq: term -> term -> term
   val mk_proj: term -> term
   val mk_quotient: term -> term -> term
+  val mk_rec_list: term -> term -> term
+  val mk_rec_nat: term -> term -> term
   val mk_shift: term -> term -> term
   val mk_size: term -> term
   val mk_toCard: term -> term -> term
@@ -146,16 +146,16 @@
 
 fun mk_undefined T = Const (@{const_name undefined}, T);
 
-fun mk_nat_rec Zero Suc =
+fun mk_rec_nat Zero Suc =
   let val T = fastype_of Zero;
-  in Const (@{const_name nat_rec}, T --> fastype_of Suc --> HOLogic.natT --> T) $ Zero $ Suc end;
+  in Const (@{const_name old.rec_nat}, T --> fastype_of Suc --> HOLogic.natT --> T) $ Zero $ Suc end;
 
-fun mk_list_rec Nil Cons =
+fun mk_rec_list Nil Cons =
   let
     val T = fastype_of Nil;
     val (U, consT) = `(Term.domain_type) (fastype_of Cons);
   in
-    Const (@{const_name list_rec}, T --> consT --> HOLogic.listT U --> T) $ Nil $ Cons
+    Const (@{const_name rec_list}, T --> consT --> HOLogic.listT U --> T) $ Nil $ Cons
   end;
 
 fun mk_InN_not_InM 1 _ = @{thm Inl_not_Inr}
--- a/src/HOL/Tools/BNF/bnf_lfp_compat.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp_compat.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -99,9 +99,9 @@
       else
         ((fp_sugars0, (NONE, NONE)), lthy);
 
-    val {ctr_sugars, co_inducts = [induct], co_iterss, co_iter_thmsss = iter_thmsss, ...} :: _ =
-      fp_sugars;
-    val inducts = conj_dests nn induct;
+    val {ctr_sugars, co_inducts = [induct], co_inductss = inductss, co_iterss,
+      co_iter_thmsss = iter_thmsss, ...} :: _ = fp_sugars;
+    val inducts = map the_single inductss;
 
     val mk_dtyp = dtyp_of_typ Ts;
 
--- a/src/HOL/Tools/BNF/bnf_lfp_tactics.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp_tactics.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -724,7 +724,7 @@
         passive_set_map0s ctor_set_incls),
         CONJ_WRAP' (fn (in_Irel, (set_map0, ctor_set_set_incls)) =>
           EVERY' [rtac ord_eq_le_trans, rtac set_map0, rtac @{thm image_subsetI}, rtac CollectI,
-            rtac @{thm prod_caseI}, rtac (in_Irel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+            rtac @{thm case_prodI}, rtac (in_Irel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
             CONJ_WRAP' (fn thm =>
               EVERY' (map etac [thm RS subset_trans, le_arg_cong_ctor_dtor]))
             ctor_set_set_incls,
@@ -750,7 +750,7 @@
                 dtac set_rev_mp, etac @{thm image_mono}, etac imageE,
                 dtac @{thm ssubst_mem[OF pair_collapse]},
                 REPEAT_DETERM o eresolve_tac (CollectE :: conjE ::
-                  @{thms prod_caseE iffD1[OF Pair_eq, elim_format]}),
+                  @{thms case_prodE iffD1[OF Pair_eq, elim_format]}),
                 hyp_subst_tac ctxt,
                 dtac (in_Irel RS iffD1), dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE,
                 TRY o
@@ -765,7 +765,7 @@
           EVERY' (map (fn in_Irel => EVERY' [rtac trans, rtac o_apply, dtac set_rev_mp, atac,
             dtac @{thm ssubst_mem[OF pair_collapse]},
             REPEAT_DETERM o
-              eresolve_tac (CollectE :: conjE :: @{thms prod_caseE iffD1[OF Pair_eq, elim_format]}),
+              eresolve_tac (CollectE :: conjE :: @{thms case_prodE iffD1[OF Pair_eq, elim_format]}),
             hyp_subst_tac ctxt,
             dtac (in_Irel RS iffD1), dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, atac])
           in_Irels),
--- a/src/HOL/Tools/Ctr_Sugar/case_translation.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Ctr_Sugar/case_translation.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -221,8 +221,9 @@
     val constr_keys = map (fst o dest_Const) constrs;
     val data = (case_comb, constrs);
     val Tname = Tname_of_data data;
-    val update_constrs = fold (fn key => Symtab.cons_list (key, (Tname, data))) constr_keys;
-    val update_cases = Symtab.update (case_key, data);
+    val update_constrs =
+      fold (fn key => Symtab.insert_list (eq_fst (op =)) (key, (Tname, data))) constr_keys;
+    val update_cases = Symtab.default (case_key, data);
   in
     context
     |> map_constrs update_constrs
--- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -41,8 +41,6 @@
   val register_ctr_sugar: string -> ctr_sugar -> local_theory -> local_theory
   val register_ctr_sugar_global: string -> ctr_sugar -> theory -> theory
 
-  val rep_compat_prefix: string
-
   val mk_half_pairss: 'a list * 'a list -> ('a * 'a) list list
   val join_halves: int -> 'a list list -> 'a list list -> 'a list * 'a list list list
 
@@ -56,10 +54,10 @@
     (ctr_sugar * term list * term list) option
 
   val wrap_free_constructors: ({prems: thm list, context: Proof.context} -> tactic) list list ->
-    (((bool * (bool * bool)) * term list) * binding) *
+    (((bool * bool) * term list) * binding) *
       (binding list * (binding list list * (binding * term) list list)) -> local_theory ->
     ctr_sugar * local_theory
-  val parse_wrap_free_constructors_options: (bool * (bool * bool)) parser
+  val parse_wrap_free_constructors_options: (bool * bool) parser
   val parse_bound_term: (binding * string) parser
 end;
 
@@ -96,10 +94,6 @@
    sel_split_asms: thm list,
    case_eq_ifs: thm list};
 
-fun eq_ctr_sugar ({ctrs = ctrs1, casex = case1, discs = discs1, selss = selss1, ...} : ctr_sugar,
-    {ctrs = ctrs2, casex = case2, discs = discs2, selss = selss2, ...} : ctr_sugar) =
-  ctrs1 = ctrs2 andalso case1 = case2 andalso discs1 = discs2 andalso selss1 = selss2;
-
 fun morph_ctr_sugar phi {ctrs, casex, discs, selss, exhaust, nchotomy, injects, distincts,
     case_thms, case_cong, weak_case_cong, split, split_asm, disc_thmss, discIs, sel_thmss,
     disc_excludesss, disc_exhausts, sel_exhausts, collapses, expands, sel_splits, sel_split_asms,
@@ -137,7 +131,7 @@
   type T = ctr_sugar Symtab.table;
   val empty = Symtab.empty;
   val extend = I;
-  val merge = Symtab.merge eq_ctr_sugar;
+  fun merge data : T = Symtab.merge (K true) data;
 );
 
 fun ctr_sugar_of ctxt =
@@ -157,8 +151,6 @@
 fun register_ctr_sugar_global key ctr_sugar =
   Context.theory_map (Data.map (Symtab.default (key, ctr_sugar)));
 
-val rep_compat_prefix = "new";
-
 val isN = "is_";
 val unN = "un_";
 fun mk_unN 1 1 suf = unN ^ suf
@@ -286,7 +278,7 @@
 
 fun eta_expand_arg xs f_xs = fold_rev Term.lambda xs f_xs;
 
-fun prepare_wrap_free_constructors prep_term ((((no_discs_sels, (no_code, rep_compat)), raw_ctrs),
+fun prepare_wrap_free_constructors prep_term ((((no_discs_sels, no_code), raw_ctrs),
     raw_case_binding), (raw_disc_bindings, (raw_sel_bindingss, raw_sel_defaultss))) no_defs_lthy =
   let
     (* TODO: sanity checks on arguments *)
@@ -304,8 +296,7 @@
     val fc_b_name = Long_Name.base_name fcT_name;
     val fc_b = Binding.name fc_b_name;
 
-    fun qualify mandatory =
-      Binding.qualify mandatory fc_b_name o (rep_compat ? Binding.qualify false rep_compat_prefix);
+    fun qualify mandatory = Binding.qualify mandatory fc_b_name;
 
     fun dest_TFree_or_TVar (TFree sS) = sS
       | dest_TFree_or_TVar (TVar ((s, _), S)) = (s, S)
@@ -363,8 +354,10 @@
 
     val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss;
 
-    val ((((((((xss, xss'), yss), fs), gs), [u', v']), [w]), (p, p')), names_lthy) = no_defs_lthy |>
-      mk_Freess' "x" ctr_Tss
+    val (((((((([exh_y], (xss, xss')), yss), fs), gs), [u', v']), [w]), (p, p')), names_lthy) =
+      no_defs_lthy
+      |> mk_Frees "y" [fcT] (* for compatibility with "datatype_realizer.ML" *)
+      ||>> mk_Freess' "x" ctr_Tss
       ||>> mk_Freess "y" ctr_Tss
       ||>> mk_Frees "f" case_Ts
       ||>> mk_Frees "g" case_Ts
@@ -443,7 +436,7 @@
 
     val (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') =
       if no_discs_sels then
-        (true, [], [], [], [], [], lthy)
+        (true, [], [], [], [], [], lthy')
       else
         let
           fun disc_free b = Free (Binding.name_of b, mk_pred1T fcT);
@@ -533,8 +526,8 @@
     fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p);
 
     val exhaust_goal =
-      let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (u, xctr)]) in
-        fold_rev Logic.all [p, u] (mk_imp_p (map2 mk_prem xctrs xss))
+      let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (exh_y, xctr)]) in
+        fold_rev Logic.all [p, exh_y] (mk_imp_p (map2 mk_prem xctrs xss))
       end;
 
     val inject_goalss =
@@ -560,7 +553,10 @@
 
     fun after_qed thmss lthy =
       let
-        val ([exhaust_thm], (inject_thmss, half_distinct_thmss)) = (hd thmss, chop n (tl thmss));
+        val ([exhaust_thm0], (inject_thmss, half_distinct_thmss)) = (hd thmss, chop n (tl thmss));
+        (* for "datatype_realizer.ML": *)
+        val exhaust_thm =
+          Thm.name_derivation (fcT_name ^ Long_Name.separator ^ exhaustN) exhaust_thm0;
 
         val inject_thms = flat inject_thmss;
 
@@ -615,7 +611,8 @@
           in
             (Goal.prove_sorry lthy [] [] goal (fn _ => mk_case_cong_tac lthy uexhaust_thm case_thms),
              Goal.prove_sorry lthy [] [] weak_goal (K (etac arg_cong 1)))
-            |> pairself (Thm.close_derivation #> singleton (Proof_Context.export names_lthy lthy))
+            |> pairself (singleton (Proof_Context.export names_lthy lthy) #>
+              Thm.close_derivation)
           end;
 
         val split_lhs = q $ ufcase;
@@ -636,14 +633,14 @@
         fun prove_split selss goal =
           Goal.prove_sorry lthy [] [] goal (fn _ =>
             mk_split_tac lthy uexhaust_thm case_thms selss inject_thmss distinct_thmsss)
-          |> Thm.close_derivation
-          |> singleton (Proof_Context.export names_lthy lthy);
+          |> singleton (Proof_Context.export names_lthy lthy)
+          |> Thm.close_derivation;
 
         fun prove_split_asm asm_goal split_thm =
           Goal.prove_sorry lthy [] [] asm_goal (fn {context = ctxt, ...} =>
             mk_split_asm_tac ctxt split_thm)
-          |> Thm.close_derivation
-          |> singleton (Proof_Context.export names_lthy lthy);
+          |> singleton (Proof_Context.export names_lthy lthy)
+          |> Thm.close_derivation;
 
         val (split_thm, split_asm_thm) =
           let
@@ -697,8 +694,8 @@
                   val goal = mk_Trueprop_eq (mk_uu_eq (), the_single exist_xs_u_eq_ctrs);
                 in
                   Goal.prove_sorry lthy [] [] goal (fn _ => mk_unique_disc_def_tac m uexhaust_thm)
+                  |> singleton (Proof_Context.export names_lthy lthy)
                   |> Thm.close_derivation
-                  |> singleton (Proof_Context.export names_lthy lthy)
                 end;
 
               fun mk_alternate_disc_def k =
@@ -710,8 +707,8 @@
                   Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
                     mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k))
                       (nth distinct_thms (2 - k)) uexhaust_thm)
+                  |> singleton (Proof_Context.export names_lthy lthy)
                   |> Thm.close_derivation
-                  |> singleton (Proof_Context.export names_lthy lthy)
                 end;
 
               val has_alternate_disc_def =
@@ -847,8 +844,8 @@
                     mk_expand_tac lthy n ms (inst_thm u disc_exhaust_thm)
                       (inst_thm v disc_exhaust_thm) uncollapse_thms disc_exclude_thmsss
                       disc_exclude_thmsss')
+                  |> singleton (Proof_Context.export names_lthy lthy)
                   |> Thm.close_derivation
-                  |> singleton (Proof_Context.export names_lthy lthy)
                 end;
 
               val (sel_split_thm, sel_split_asm_thm) =
@@ -869,8 +866,8 @@
                 in
                   Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
                     mk_case_eq_if_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss)
+                  |> singleton (Proof_Context.export names_lthy lthy)
                   |> Thm.close_derivation
-                  |> singleton (Proof_Context.export names_lthy lthy)
                 end;
             in
               (all_sel_thms, sel_thmss, disc_thmss, nontriv_disc_thms, discI_thms,
@@ -927,8 +924,7 @@
       in
         (ctr_sugar,
          lthy
-         |> not rep_compat ?
-            Local_Theory.declaration {syntax = false, pervasive = true}
+         |> Local_Theory.declaration {syntax = false, pervasive = true}
               (fn phi => Case_Translation.register
                  (Morphism.term phi casex) (map (Morphism.term phi) ctrs))
          |> Local_Theory.background_theory (fold (fold Code.del_eqn) [disc_defs, sel_defs])
@@ -965,10 +961,10 @@
 
 val parse_wrap_free_constructors_options =
   Scan.optional (@{keyword "("} |-- Parse.list1
-        (Parse.reserved "no_discs_sels" >> K 0 || Parse.reserved "no_code" >> K 1 ||
-         Parse.reserved "rep_compat" >> K 2) --| @{keyword ")"}
-      >> (fn js => (member (op =) js 0, (member (op =) js 1, member (op =) js 2))))
-    (false, (false, false));
+        (Parse.reserved "no_discs_sels" >> K 0 || Parse.reserved "no_code" >> K 1) --|
+      @{keyword ")"}
+      >> (fn js => (member (op =) js 0, member (op =) js 1)))
+    (false, false);
 
 val _ =
   Outer_Syntax.local_theory_to_proof @{command_spec "wrap_free_constructors"}
--- a/src/HOL/Tools/Datatype/datatype.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Datatype/datatype.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -62,7 +62,7 @@
     val new_type_names = map (Binding.name_of o fst) types_syntax;
     val big_name = space_implode "_" new_type_names;
     val thy1 = Sign.add_path big_name thy;
-    val big_rec_name = big_name ^ "_rep_set";
+    val big_rec_name = "rep_set_" ^ big_name;
     val rep_set_names' =
       if length descr' = 1 then [big_rec_name]
       else map (prefix (big_rec_name ^ "_") o string_of_int) (1 upto length descr');
@@ -283,11 +283,11 @@
     (* isomorphisms are defined using primrec-combinators:                 *)
     (* generate appropriate functions for instantiating primrec-combinator *)
     (*                                                                     *)
-    (*   e.g.  dt_Rep_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y))    *)
+    (*   e.g.  Rep_dt_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y))    *)
     (*                                                                     *)
     (* also generate characteristic equations for isomorphisms             *)
     (*                                                                     *)
-    (*   e.g.  dt_Rep_i (cons h t) = In1 (Scons (dt_Rep_j h) (dt_Rep_i t)) *)
+    (*   e.g.  Rep_dt_i (cons h t) = In1 (Scons (Rep_dt_j h) (Rep_dt_i t)) *)
     (*---------------------------------------------------------------------*)
 
     fun make_iso_def k ks n (cname, cargs) (fs, eqns, i) =
@@ -387,7 +387,7 @@
           end
       in map (fn r => r RS subst) (thm :: map mk_thm arities) end;
 
-    (* prove  inj dt_Rep_i  and  dt_Rep_i x : dt_rep_set_i *)
+    (* prove  inj Rep_dt_i  and  Rep_dt_i x : rep_set_dt_i *)
 
     val fun_congs =
       map (fn T => make_elim (Drule.instantiate' [SOME (ctyp_of thy5 T)] [] fun_cong)) branchTs;
@@ -457,7 +457,7 @@
     val iso_inj_thms =
       map snd newT_iso_inj_thms @ map (fn r => r RS @{thm injD}) iso_inj_thms_unfolded;
 
-    (* prove  dt_rep_set_i x --> x : range dt_Rep_i *)
+    (* prove  rep_set_dt_i x --> x : range Rep_dt_i *)
 
     fun mk_iso_t (((set_name, iso_name), i), T) =
       let val isoT = T --> Univ_elT in
--- a/src/HOL/Tools/Datatype/datatype_codegen.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Datatype/datatype_codegen.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -13,15 +13,11 @@
 
 fun add_code_for_datatype fcT_name thy =
   let
-    val (As', ctr_specs) = Datatype_Data.the_spec thy fcT_name;
-    val {inject = inject_thms, distinct = distinct_thms, case_rewrites = case_thms, ...} =
-      Datatype_Data.the_info thy fcT_name;
-
-    val As = map TFree As';
-    val fcT = Type (fcT_name, As);
-    val ctrs = map (fn (c, arg_Ts) => (c, arg_Ts ---> fcT)) ctr_specs;
+    val ctxt = Proof_Context.init_global thy
+    val SOME {ctrs, injects, distincts, case_thms, ...} = Ctr_Sugar.ctr_sugar_of ctxt fcT_name
+    val Type (_, As) = body_type (fastype_of (hd ctrs))
   in
-    Ctr_Sugar_Code.add_ctr_code fcT_name As ctrs inject_thms distinct_thms case_thms thy
+    Ctr_Sugar_Code.add_ctr_code fcT_name As (map dest_Const ctrs) injects distincts case_thms thy
   end;
 
 val _ = Theory.setup (Datatype_Data.interpretation (K (fold add_code_for_datatype)));
--- a/src/HOL/Tools/Datatype/rep_datatype.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Datatype/rep_datatype.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -87,7 +87,7 @@
 
     val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
 
-    val big_rec_name' = big_name ^ "_rec_set";
+    val big_rec_name' = "rec_set_" ^ big_name;
     val rec_set_names' =
       if length descr' = 1 then [big_rec_name']
       else map (prefix (big_rec_name' ^ "_") o string_of_int) (1 upto length descr');
@@ -215,7 +215,7 @@
 
     (* define primrec combinators *)
 
-    val big_reccomb_name = space_implode "_" new_type_names ^ "_rec";
+    val big_reccomb_name = "rec_" ^ space_implode "_" new_type_names;
     val reccomb_names =
       map (Sign.full_bname thy1)
         (if length descr' = 1 then [big_reccomb_name]
@@ -271,6 +271,7 @@
   let
     val _ = Datatype_Aux.message config "Proving characteristic theorems for case combinators ...";
 
+    val ctxt = Proof_Context.init_global thy;
     val thy1 = Sign.add_path (space_implode "_" new_type_names) thy;
 
     val descr' = flat descr;
@@ -288,48 +289,62 @@
           val Ts' = map mk_dummyT (filter Datatype_Aux.is_rec_type cargs)
         in Const (@{const_name undefined}, Ts @ Ts' ---> T') end) constrs) descr';
 
-    val case_names = map (fn s => Sign.full_bname thy1 (s ^ "_case")) new_type_names;
+    val case_names0 = map (fn s => Sign.full_bname thy1 ("case_" ^ s)) new_type_names;
 
     (* define case combinators via primrec combinators *)
 
-    val (case_defs, thy2) =
-      fold (fn ((((i, (_, _, constrs)), T), name), recname) => fn (defs, thy) =>
-          let
-            val (fns1, fns2) = split_list (map (fn ((_, cargs), j) =>
-              let
-                val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
-                val Ts' = Ts @ map mk_dummyT (filter Datatype_Aux.is_rec_type cargs);
-                val frees' = map2 (Datatype_Aux.mk_Free "x") Ts' (1 upto length Ts');
-                val frees = take (length cargs) frees';
-                val free = Datatype_Aux.mk_Free "f" (Ts ---> T') j;
-              in
-                (free, fold_rev (absfree o dest_Free) frees' (list_comb (free, frees)))
-              end) (constrs ~~ (1 upto length constrs)));
+    fun def_case ((((i, (_, _, constrs)), T as Type (Tcon, _)), name), recname) (defs, thy) =
+      if is_some (Ctr_Sugar.ctr_sugar_of ctxt Tcon) then
+        (defs, thy)
+      else
+        let
+          val (fns1, fns2) = split_list (map (fn ((_, cargs), j) =>
+            let
+              val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
+              val Ts' = Ts @ map mk_dummyT (filter Datatype_Aux.is_rec_type cargs);
+              val frees' = map2 (Datatype_Aux.mk_Free "x") Ts' (1 upto length Ts');
+              val frees = take (length cargs) frees';
+              val free = Datatype_Aux.mk_Free "f" (Ts ---> T') j;
+            in
+              (free, fold_rev (absfree o dest_Free) frees' (list_comb (free, frees)))
+            end) (constrs ~~ (1 upto length constrs)));
 
-            val caseT = map (snd o dest_Free) fns1 @ [T] ---> T';
-            val fns = flat (take i case_dummy_fns) @ fns2 @ flat (drop (i + 1) case_dummy_fns);
-            val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T');
-            val decl = ((Binding.name (Long_Name.base_name name), caseT), NoSyn);
-            val def =
-              (Binding.name (Thm.def_name (Long_Name.base_name name)),
-                Logic.mk_equals (Const (name, caseT),
-                  fold_rev lambda fns1
-                    (list_comb (reccomb,
-                      flat (take i case_dummy_fns) @ fns2 @ flat (drop (i + 1) case_dummy_fns)))));
-            val ([def_thm], thy') =
-              thy
-              |> Sign.declare_const_global decl |> snd
-              |> (Global_Theory.add_defs false o map Thm.no_attributes) [def];
+          val caseT = map (snd o dest_Free) fns1 @ [T] ---> T';
+          val fns = flat (take i case_dummy_fns) @ fns2 @ flat (drop (i + 1) case_dummy_fns);
+          val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T');
+          val decl = ((Binding.name (Long_Name.base_name name), caseT), NoSyn);
+          val def =
+            (Binding.name (Thm.def_name (Long_Name.base_name name)),
+              Logic.mk_equals (Const (name, caseT),
+                fold_rev lambda fns1
+                  (list_comb (reccomb,
+                    flat (take i case_dummy_fns) @ fns2 @ flat (drop (i + 1) case_dummy_fns)))));
+          val ([def_thm], thy') =
+            thy
+            |> Sign.declare_const_global decl |> snd
+            |> (Global_Theory.add_defs false o map Thm.no_attributes) [def];
+        in (defs @ [def_thm], thy') end;
 
-          in (defs @ [def_thm], thy') end)
-        (hd descr ~~ newTs ~~ case_names ~~ take (length newTs) reccomb_names) ([], thy1);
+    val (case_defs, thy2) =
+      fold def_case (hd descr ~~ newTs ~~ case_names0 ~~ take (length newTs) reccomb_names)
+        ([], thy1);
+
+    fun prove_case t =
+      Goal.prove_sorry_global thy2 [] [] t (fn {context = ctxt, ...} =>
+        EVERY [rewrite_goals_tac ctxt (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1]);
+
+    fun prove_cases (Type (Tcon, _)) ts =
+      (case Ctr_Sugar.ctr_sugar_of ctxt Tcon of
+        SOME {case_thms, ...} => case_thms
+      | NONE => map prove_case ts);
 
     val case_thms =
-      (map o map) (fn t =>
-          Goal.prove_sorry_global thy2 [] [] t
-            (fn {context = ctxt, ...} =>
-              EVERY [rewrite_goals_tac ctxt (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1]))
-        (Datatype_Prop.make_cases case_names descr thy2);
+      map2 prove_cases newTs (Datatype_Prop.make_cases case_names0 descr thy2);
+
+    fun case_name_of (th :: _) =
+      fst (dest_Const (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of th))))));
+
+    val case_names = map case_name_of case_thms;
   in
     thy2
     |> Context.theory_map ((fold o fold) Nitpick_Simps.add_thm case_thms)
--- a/src/HOL/Tools/Function/function.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Function/function.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -265,14 +265,14 @@
 
 fun add_case_cong n thy =
   let
-    val cong = #case_cong (Datatype.the_info thy n)
+    val cong = #case_cong (Datatype_Data.the_info thy n)
       |> safe_mk_meta_eq
   in
     Context.theory_map
       (Function_Ctx_Tree.map_function_congs (Thm.add_thm cong)) thy
   end
 
-val setup_case_cong = Datatype.interpretation (K (fold add_case_cong))
+val setup_case_cong = Datatype_Data.interpretation (K (fold add_case_cong))
 
 
 (* setup *)
--- a/src/HOL/Tools/Function/function_core.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Function/function_core.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -829,7 +829,7 @@
   let
     val FunctionConfig {domintros, default=default_opt, ...} = config
 
-    val default_str = the_default "%x. undefined" default_opt (*FIXME dynamic scoping*)
+    val default_str = the_default "%x. HOL.undefined" default_opt
     val fvar = Free (fname, fT)
     val domT = domain_type fT
     val ranT = range_type fT
--- a/src/HOL/Tools/Function/sum_tree.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Function/sum_tree.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -32,8 +32,7 @@
 (* Sum types *)
 fun mk_sumT LT RT = Type (@{type_name Sum_Type.sum}, [LT, RT])
 fun mk_sumcase TL TR T l r =
-  Const (@{const_name sum.sum_case},
-    (TL --> T) --> (TR --> T) --> mk_sumT TL TR --> T) $ l $ r
+  Const (@{const_name sum.case_sum}, (TL --> T) --> (TR --> T) --> mk_sumT TL TR --> T) $ l $ r
 
 val App = curry op $
 
@@ -50,9 +49,9 @@
   access_top_down
   { init = (ST, I : term -> term),
     left = (fn (T as Type (@{type_name Sum_Type.sum}, [LT, RT]), proj) =>
-      (LT, App (Const (@{const_name Sum_Type.Projl}, T --> LT)) o proj)),
+      (LT, App (Const (@{const_name Sum_Type.projl}, T --> LT)) o proj)),
     right =(fn (T as Type (@{type_name Sum_Type.sum}, [LT, RT]), proj) =>
-      (RT, App (Const (@{const_name Sum_Type.Projr}, T --> RT)) o proj))} n i
+      (RT, App (Const (@{const_name Sum_Type.projr}, T --> RT)) o proj))} n i
   |> snd
 
 fun mk_sumcases T fs =
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -2267,7 +2267,7 @@
       HOLogic.Collect_const tuple_T $ list_comb (Const base_x, outer_bounds)
     val step_set =
       HOLogic.Collect_const prod_T
-      $ (Const (@{const_name prod_case}, curried_T --> uncurried_T)
+      $ (Const (@{const_name case_prod}, curried_T --> uncurried_T)
                 $ list_comb (Const step_x, outer_bounds))
     val image_set =
       image_const $ (rtrancl_const $ step_set) $ base_set
--- a/src/HOL/Tools/Predicate_Compile/code_prolog.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Predicate_Compile/code_prolog.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -511,13 +511,6 @@
 
 fun mk_lim_relname T = "lim_" ^  mk_relname T
 
-(* This is copied from "pat_completeness.ML" *)
-fun inst_constrs_of thy (T as Type (name, _)) =
-  map (fn (Cn,CT) =>
-    Envir.subst_term_types (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT)))
-    (the (Datatype.get_constrs thy name))
-  | inst_constrs_of thy T = raise TYPE ("inst_constrs_of", [T], [])
-
 fun is_recursive_constr T (Const (constr_name, T')) = member (op =) (binder_types T') T
   
 fun mk_ground_impl ctxt limited_types (T as Type (Tcon, Targs)) (seen, constant_table) =
@@ -549,7 +542,7 @@
         in
           (clause :: flat rec_clauses, (seen', constant_table''))
         end
-      val constrs = inst_constrs_of (Proof_Context.theory_of ctxt) T
+      val constrs = Function_Lib.inst_constrs_of (Proof_Context.theory_of ctxt) T
       val constrs' = (constrs ~~ map (is_recursive_constr T) constrs)
         |> (fn cs => filter_out snd cs @ filter snd cs)
       val (clauses, constant_table') =
--- a/src/HOL/Tools/Predicate_Compile/mode_inference.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Predicate_Compile/mode_inference.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -183,19 +183,19 @@
   | collect_non_invertible_subterms ctxt t (names, eqs) =
     case (strip_comb t) of (f, args) =>
       if is_invertible_function ctxt f then
-          let
-            val (args', (names', eqs')) =
-              fold_map (collect_non_invertible_subterms ctxt) args (names, eqs)
-          in
-            (list_comb (f, args'), (names', eqs'))
-          end
-        else
-          let
-            val s = singleton (Name.variant_list names) "x"
-            val v = Free (s, fastype_of t)
-          in
-            (v, (s :: names, HOLogic.mk_eq (v, t) :: eqs))
-          end
+        let
+          val (args', (names', eqs')) =
+            fold_map (collect_non_invertible_subterms ctxt) args (names, eqs)
+        in
+          (list_comb (f, args'), (names', eqs'))
+        end
+      else
+        let
+          val s = singleton (Name.variant_list names) "x"
+          val v = Free (s, fastype_of t)
+        in
+          (v, (s :: names, HOLogic.mk_eq (v, t) :: eqs))
+        end
 (*
   if is_constrt thy t then (t, (names, eqs)) else
     let
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -47,6 +47,7 @@
   val is_pred_equation : thm -> bool
   val is_intro : string -> thm -> bool
   val is_predT : typ -> bool
+  val get_constrs : theory -> (string * (int * string)) list
   val is_constrt : theory -> term -> bool
   val is_constr : Proof.context -> string -> bool
   val strip_ex : term -> (string * typ) list * term
@@ -477,15 +478,22 @@
 fun is_predT (T as Type("fun", [_, _])) = (body_type T = @{typ bool})
   | is_predT _ = false
 
+fun get_constrs thy =
+  let
+    val ctxt = Proof_Context.init_global thy
+  in
+    Ctr_Sugar.ctr_sugars_of ctxt
+    |> maps (map_filter (try dest_Const) o #ctrs)
+    |> map (apsnd (fn T => (BNF_Util.num_binder_types T, fst (dest_Type (body_type T)))))
+  end
+
 (*** check if a term contains only constructor functions ***)
 (* TODO: another copy in the core! *)
 (* FIXME: constructor terms are supposed to be seen in the way the code generator
   sees constructors.*)
 fun is_constrt thy =
   let
-    val cnstrs = flat (maps
-      (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
-      (Symtab.dest (Datatype.get_all thy)));
+    val cnstrs = get_constrs thy
     fun check t = (case strip_comb t of
         (Var _, []) => true
       | (Free _, []) => true
@@ -495,23 +503,6 @@
       | _ => false)
   in check end;
 
-(* returns true if t is an application of an datatype constructor *)
-(* which then consequently would be splitted *)
-(* else false *)
-(*
-fun is_constructor thy t =
-  if (is_Type (fastype_of t)) then
-    (case DatatypePackage.get_datatype thy ((fst o dest_Type o fastype_of) t) of
-      NONE => false
-    | SOME info => (let
-      val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info)
-      val (c, _) = strip_comb t
-      in (case c of
-        Const (name, _) => name mem_string constr_consts
-        | _ => false) end))
-  else false
-*)
-
 val is_constr = Code.is_constr o Proof_Context.theory_of;
 
 fun strip_all t = (Term.strip_all_vars t, Term.strip_all_body t)
@@ -601,7 +592,8 @@
     |> Local_Defs.unfold ctxt [@{thm atomize_conjL[symmetric]},
       @{thm atomize_all[symmetric]}, @{thm atomize_imp[symmetric]}]
 
-fun find_split_thm thy (Const (name, _)) = Option.map #split (Datatype.info_of_case thy name)
+fun find_split_thm thy (Const (name, _)) =
+    Option.map #split (Ctr_Sugar.ctr_sugar_of_case (Proof_Context.init_global thy) name)
   | find_split_thm thy _ = NONE
 
 (* lifting term operations to theorems *)
@@ -880,76 +872,72 @@
 (** making case distributivity rules **)
 (*** this should be part of the datatype package ***)
 
-fun datatype_names_of_case_name thy case_name =
-  map (#1 o #2) (#descr (the (Datatype.info_of_case thy case_name)))
+fun datatype_name_of_case_name thy =
+  Ctr_Sugar.ctr_sugar_of_case (Proof_Context.init_global thy)
+  #> the #> #ctrs #> hd #> fastype_of #> body_type #> dest_Type #> fst
 
-fun make_case_distribs case_names descr thy =
+fun make_case_comb thy Tcon =
   let
-    val case_combs = Datatype_Prop.make_case_combs case_names descr thy "f";
-    fun make comb =
-      let
-        val Type ("fun", [T, T']) = fastype_of comb;
-        val (Const (case_name, _), fs) = strip_comb comb
-        val used = Term.add_tfree_names comb []
-        val U = TFree (singleton (Name.variant_list used) "'t", HOLogic.typeS)
-        val x = Free ("x", T)
-        val f = Free ("f", T' --> U)
-        fun apply_f f' =
-          let
-            val Ts = binder_types (fastype_of f')
-            val bs = map Bound ((length Ts - 1) downto 0)
-          in
-            fold_rev absdummy Ts (f $ (list_comb (f', bs)))
-          end
-        val fs' = map apply_f fs
-        val case_c' = Const (case_name, (map fastype_of fs') @ [T] ---> U)
-      in
-        HOLogic.mk_eq (f $ (comb $ x), list_comb (case_c', fs') $ x)
-      end
+    val ctxt = Proof_Context.init_global thy
+    val SOME {casex, ...} = Ctr_Sugar.ctr_sugar_of ctxt Tcon
+    val casex' = Type.legacy_freeze casex
+    val Ts = BNF_Util.binder_fun_types (fastype_of casex')
   in
-    map make case_combs
+    list_comb (casex', map_index (fn (j, T) => Free ("f" ^ string_of_int j,  T)) Ts)
   end
 
-fun case_rewrites thy Tcon =
+fun make_case_distrib thy Tcon =
   let
-    val {descr, case_name, ...} = Datatype.the_info thy Tcon
+    val comb = make_case_comb thy Tcon;
+    val Type ("fun", [T, T']) = fastype_of comb;
+    val (Const (case_name, _), fs) = strip_comb comb
+    val used = Term.add_tfree_names comb []
+    val U = TFree (singleton (Name.variant_list used) "'t", HOLogic.typeS)
+    val x = Free ("x", T)
+    val f = Free ("f", T' --> U)
+    fun apply_f f' =
+      let
+        val Ts = binder_types (fastype_of f')
+        val bs = map Bound ((length Ts - 1) downto 0)
+      in
+        fold_rev absdummy Ts (f $ (list_comb (f', bs)))
+      end
+    val fs' = map apply_f fs
+    val case_c' = Const (case_name, (map fastype_of fs') @ [T] ---> U)
   in
-    map (Drule.export_without_context o Skip_Proof.make_thm thy o HOLogic.mk_Trueprop)
-      (make_case_distribs [case_name] [descr] thy)
+    HOLogic.mk_eq (f $ (comb $ x), list_comb (case_c', fs') $ x)
   end
 
-fun instantiated_case_rewrites thy Tcon =
+fun case_rewrite thy Tcon =
+  (Drule.export_without_context o Skip_Proof.make_thm thy o HOLogic.mk_Trueprop)
+    (make_case_distrib thy Tcon)
+
+fun instantiated_case_rewrite thy Tcon =
   let
-    val rew_ths = case_rewrites thy Tcon
+    val th = case_rewrite thy Tcon
     val ctxt = Proof_Context.init_global thy
-    fun instantiate th =
-    let
-      val f = (fst (strip_comb (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of th))))))
-      val Type ("fun", [uninst_T, uninst_T']) = fastype_of f
-      val ([_, tname', uname, yname], ctxt') = Variable.add_fixes ["'t", "'t'", "'u", "y"] ctxt
-      val T' = TFree (tname', HOLogic.typeS)
-      val U = TFree (uname, HOLogic.typeS)
-      val y = Free (yname, U)
-      val f' = absdummy (U --> T') (Bound 0 $ y)
-      val th' = Thm.certify_instantiate
-        ([(dest_TVar uninst_T, U --> T'), (dest_TVar uninst_T', T')],
-         [((fst (dest_Var f), (U --> T') --> T'), f')]) th
-      val [th'] = Variable.export ctxt' ctxt [th']
-   in
-     th'
-   end
- in
-   map instantiate rew_ths
- end
+    val f = (fst (strip_comb (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of th))))))
+    val Type ("fun", [uninst_T, uninst_T']) = fastype_of f
+    val ([_, tname', uname, yname], ctxt') = Variable.add_fixes ["'t", "'t'", "'u", "y"] ctxt
+    val T' = TFree (tname', HOLogic.typeS)
+    val U = TFree (uname, HOLogic.typeS)
+    val y = Free (yname, U)
+    val f' = absdummy (U --> T') (Bound 0 $ y)
+    val th' = Thm.certify_instantiate
+      ([(dest_TVar uninst_T, U --> T'), (dest_TVar uninst_T', T')],
+       [((fst (dest_Var f), (U --> T') --> T'), f')]) th
+    val [th'] = Variable.export ctxt' ctxt [th']
+  in
+    th'
+  end
 
 fun case_betapply thy t =
   let
     val case_name = fst (dest_Const (fst (strip_comb t)))
-    val Tcons = datatype_names_of_case_name thy case_name
-    val ths = maps (instantiated_case_rewrites thy) Tcons
+    val Tcon = datatype_name_of_case_name thy case_name
+    val th = instantiated_case_rewrite thy Tcon
   in
-    Raw_Simplifier.rewrite_term thy
-      (map (fn th => th RS @{thm eq_reflection}) ths) [] t
+    Raw_Simplifier.rewrite_term thy [th RS @{thm eq_reflection}] [] t
   end
 
 (*** conversions ***)
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -814,13 +814,14 @@
     case T of
       TFree _ => NONE
     | Type (Tcon, _) =>
-      (case Datatype.get_constrs (Proof_Context.theory_of ctxt) Tcon of
+      (case Ctr_Sugar.ctr_sugar_of ctxt Tcon of
         NONE => NONE
-      | SOME cs =>
+      | SOME {ctrs, ...} =>
         (case strip_comb t of
           (Var _, []) => NONE
         | (Free _, []) => NONE
-        | (Const (c, T), _) => if AList.defined (op =) cs c then SOME (c, T) else NONE))
+        | (Const (c, T), _) =>
+          if AList.defined (op =) (map_filter (try dest_Const) ctrs) c then SOME (c, T) else NONE))
   end
 
 fun partition_clause ctxt pos moded_clauses =
@@ -991,7 +992,7 @@
 
 (* Definition of executable functions and their intro and elim rules *)
 
-fun strip_split_abs (Const (@{const_name prod_case}, _) $ t) = strip_split_abs t
+fun strip_split_abs (Const (@{const_name case_prod}, _) $ t) = strip_split_abs t
   | strip_split_abs (Abs (_, _, t)) = strip_split_abs t
   | strip_split_abs t = t
 
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -260,8 +260,11 @@
     val ctxt = Proof_Context.init_global thy
     fun is_nondefining_const (c, _) = member (op =) logic_operator_names c
     fun has_code_pred_intros (c, _) = can (Core_Data.intros_of ctxt) c
-    fun case_consts (c, _) = is_some (Datatype.info_of_case thy c)
-    fun is_datatype_constructor (c, T) = is_some (Datatype.info_of_constr thy (c, T))
+    fun case_consts (c, _) = is_some (Ctr_Sugar.ctr_sugar_of_case ctxt c)
+    fun is_datatype_constructor (x as (_, T)) =
+      (case body_type T of
+        Type (Tcon, _) => can (Ctr_Sugar.dest_ctr ctxt Tcon) (Const x)
+      | _ => false)
     fun defiants_of specs =
       fold (Term.add_consts o prop_of) specs []
       |> filter_out is_datatype_constructor
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_proof.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_proof.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -46,30 +46,19 @@
 
 (* auxillary functions *)
 
-fun is_Type (Type _) = true
-  | is_Type _ = false
-
-(* returns true if t is an application of an datatype constructor *)
+(* returns true if t is an application of a datatype constructor *)
 (* which then consequently would be splitted *)
-(* else false *)
-fun is_constructor thy t =
-  if (is_Type (fastype_of t)) then
-    (case Datatype.get_info thy ((fst o dest_Type o fastype_of) t) of
-      NONE => false
-    | SOME info => (let
-      val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info)
-      val (c, _) = strip_comb t
-      in (case c of
-        Const (name, _) => member (op =) constr_consts name
-        | _ => false) end))
-  else false
+fun is_constructor ctxt t =
+  (case fastype_of t of
+    Type (s, _) => s <> @{type_name fun} andalso can (Ctr_Sugar.dest_ctr ctxt s) t
+  | _ => false);
 
 (* MAJOR FIXME:  prove_params should be simple
  - different form of introrule for parameters ? *)
 
 fun prove_param options ctxt nargs t deriv =
   let
-    val  (f, args) = strip_comb (Envir.eta_contract t)
+    val (f, args) = strip_comb (Envir.eta_contract t)
     val mode = head_mode_of deriv
     val param_derivations = param_derivations_of deriv
     val ho_args = ho_args_of mode args
@@ -139,15 +128,14 @@
 
 fun prove_match options ctxt nargs out_ts =
   let
-    val thy = Proof_Context.theory_of ctxt
     val eval_if_P =
       @{lemma "P ==> Predicate.eval x z ==> Predicate.eval (if P then x else y) z" by simp} 
     fun get_case_rewrite t =
-      if (is_constructor thy t) then
+      if is_constructor ctxt t then
         let
-          val {case_rewrites, ...} = Datatype.the_info thy (fst (dest_Type (fastype_of t)))
+          val SOME {case_thms, ...} = Ctr_Sugar.ctr_sugar_of ctxt (fst (dest_Type (fastype_of t)))
         in
-          fold (union Thm.eq_thm) (case_rewrites :: map get_case_rewrite (snd (strip_comb t))) []
+          fold (union Thm.eq_thm) (case_thms :: map get_case_rewrite (snd (strip_comb t))) []
         end
       else []
     val simprules = insert Thm.eq_thm @{thm "unit.cases"} (insert Thm.eq_thm @{thm "prod.cases"}
@@ -309,18 +297,17 @@
 
 fun prove_match2 options ctxt out_ts =
   let
-    val thy = Proof_Context.theory_of ctxt
     fun split_term_tac (Free _) = all_tac
       | split_term_tac t =
-        if (is_constructor thy t) then
+        if is_constructor ctxt t then
           let
-            val {case_rewrites, split_asm, ...} =
-              Datatype.the_info thy (fst (dest_Type (fastype_of t)))
-            val num_of_constrs = length case_rewrites
+            val SOME {case_thms, split_asm, ...} =
+              Ctr_Sugar.ctr_sugar_of ctxt (fst (dest_Type (fastype_of t)))
+            val num_of_constrs = length case_thms
             val (_, ts) = strip_comb t
           in
             print_tac options ("Term " ^ (Syntax.string_of_term ctxt t) ^ 
-              "splitting with rules \n" ^ Display.string_of_thm ctxt split_asm)
+              " splitting with rules \n" ^ Display.string_of_thm ctxt split_asm)
             THEN TRY (Splitter.split_asm_tac [split_asm] 1
               THEN (print_tac options "after splitting with split_asm rules")
             (* THEN (Simplifier.asm_full_simp_tac (put_simpset HOL_basic_ss ctxt) 1)
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_specialisation.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_specialisation.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -41,9 +41,7 @@
 (* patterns only constructed of variables and pairs/tuples are trivial constructor terms*)
 fun is_nontrivial_constrt thy t =
   let
-    val cnstrs = flat (maps
-      (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
-      (Symtab.dest (Datatype.get_all thy)));
+    val cnstrs = get_constrs thy
     fun check t = (case strip_comb t of
         (Var _, []) => (true, true)
       | (Free _, []) => (true, true)
@@ -107,6 +105,7 @@
 
 and find_specialisations black_list specs thy =
   let
+    val ctxt = Proof_Context.init_global thy
     val add_vars = fold_aterms (fn Var v => cons v | _ => I);
     fun fresh_free T free_names =
       let
@@ -132,10 +131,11 @@
       | restrict_pattern' thy ((T as TFree _, t) :: Tts) free_names =
         replace_term_and_restrict thy T t Tts free_names
       | restrict_pattern' thy ((T as Type (Tcon, _), t) :: Tts) free_names =
-        case Datatype.get_constrs thy Tcon of
+        case Ctr_Sugar.ctr_sugar_of ctxt Tcon of
           NONE => replace_term_and_restrict thy T t Tts free_names
-        | SOME constrs => (case strip_comb t of
-          (Const (s, _), ats) => (case AList.lookup (op =) constrs s of
+        | SOME {ctrs, ...} => (case strip_comb t of
+          (Const (s, _), ats) =>
+          (case AList.lookup (op =) (map_filter (try dest_Const) ctrs) s of
             SOME constr_T =>
               let
                 val (Ts', T') = strip_type constr_T
--- a/src/HOL/Tools/Quickcheck/random_generators.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Quickcheck/random_generators.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -312,7 +312,7 @@
     fun mk_scomp T1 T2 sT f g = Const (@{const_name scomp},
       liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g;
     fun mk_split T = Sign.mk_const thy
-      (@{const_name prod_case}, [T, @{typ "unit => term"}, liftT resultT @{typ Random.seed}]);
+      (@{const_name case_prod}, [T, @{typ "unit => term"}, liftT resultT @{typ Random.seed}]);
     fun mk_scomp_split T t t' =
       mk_scomp (mk_termtyp T) resultT @{typ Random.seed} t
         (mk_split T $ Abs ("", T, Abs ("", @{typ "unit => term"}, t')));
@@ -358,7 +358,7 @@
     fun mk_scomp T1 T2 sT f g = Const (@{const_name scomp},
       liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g;
     fun mk_split T = Sign.mk_const thy
-      (@{const_name prod_case}, [T, @{typ "unit => term"}, liftT resultT @{typ Random.seed}]);
+      (@{const_name case_prod}, [T, @{typ "unit => term"}, liftT resultT @{typ Random.seed}]);
     fun mk_scomp_split T t t' =
       mk_scomp (mk_termtyp T) resultT @{typ Random.seed} t
         (mk_split T $ Abs ("", T, Abs ("", @{typ "unit => term"}, t')));
--- a/src/HOL/Tools/Quotient/quotient_term.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/Quotient/quotient_term.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -635,12 +635,12 @@
           end
       end
 
-  | (((t1 as Const (@{const_name prod_case}, _)) $ Abs (v1, ty, Abs(v1', ty', s1))),
-     ((t2 as Const (@{const_name prod_case}, _)) $ Abs (v2, _ , Abs(v2', _  , s2)))) =>
+  | (((t1 as Const (@{const_name case_prod}, _)) $ Abs (v1, ty, Abs(v1', ty', s1))),
+     ((t2 as Const (@{const_name case_prod}, _)) $ Abs (v2, _ , Abs(v2', _  , s2)))) =>
        regularize_trm ctxt (t1, t2) $ Abs (v1, ty, Abs (v1', ty', regularize_trm ctxt (s1, s2)))
 
-  | (((t1 as Const (@{const_name prod_case}, _)) $ Abs (v1, ty, s1)),
-     ((t2 as Const (@{const_name prod_case}, _)) $ Abs (v2, _ , s2))) =>
+  | (((t1 as Const (@{const_name case_prod}, _)) $ Abs (v1, ty, s1)),
+     ((t2 as Const (@{const_name case_prod}, _)) $ Abs (v2, _ , s2))) =>
        regularize_trm ctxt (t1, t2) $ Abs (v1, ty, regularize_trm ctxt (s1, s2))
 
   | (t1 $ t2, t1' $ t2') =>
--- a/src/HOL/Tools/SMT/smt_normalize.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/SMT/smt_normalize.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -356,22 +356,22 @@
 (** rewrite bool case expressions as if expressions **)
 
 local
-  fun is_bool_case (Const (@{const_name "bool.bool_case"}, _)) = true
-    | is_bool_case _ = false
+  fun is_case_bool (Const (@{const_name "bool.case_bool"}, _)) = true
+    | is_case_bool _ = false
 
   val thm = mk_meta_eq @{lemma
-    "bool_case = (%x y P. if P then x else y)" by (rule ext)+ simp}
+    "case_bool = (%x y P. if P then x else y)" by (rule ext)+ simp}
 
   fun unfold_conv _ =
-    SMT_Utils.if_true_conv (is_bool_case o Term.head_of)
+    SMT_Utils.if_true_conv (is_case_bool o Term.head_of)
       (expand_head_conv (Conv.rewr_conv thm))
 in
 
-fun rewrite_bool_case_conv ctxt =
-  SMT_Utils.if_exists_conv is_bool_case (Conv.top_conv unfold_conv ctxt)
+fun rewrite_case_bool_conv ctxt =
+  SMT_Utils.if_exists_conv is_case_bool (Conv.top_conv unfold_conv ctxt)
 
-val setup_bool_case =
-  SMT_Builtin.add_builtin_fun_ext'' @{const_name "bool.bool_case"}
+val setup_case_bool =
+  SMT_Builtin.add_builtin_fun_ext'' @{const_name "bool.case_bool"}
 
 end
 
@@ -558,7 +558,7 @@
 (** combined unfoldings and rewritings **)
 
 fun unfold_conv ctxt =
-  rewrite_bool_case_conv ctxt then_conv
+  rewrite_case_bool_conv ctxt then_conv
   unfold_abs_min_max_conv ctxt then_conv
   nat_as_int_conv ctxt then_conv
   Thm.beta_conversion true
@@ -645,7 +645,7 @@
   setup_unfolded_quants #>
   setup_trigger #>
   setup_weight #>
-  setup_bool_case #>
+  setup_case_bool #>
   setup_abs_min_max #>
   setup_nat_as_int)
 
--- a/src/HOL/Tools/TFL/rules.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/TFL/rules.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -578,10 +578,10 @@
 local fun dest_pair M = let val {fst,snd} = USyntax.dest_pair M in (fst,snd) end
       fun mk_fst tm =
           let val ty as Type(@{type_name Product_Type.prod}, [fty,sty]) = type_of tm
-          in  Const ("Product_Type.fst", ty --> fty) $ tm  end
+          in  Const (@{const_name Product_Type.fst}, ty --> fty) $ tm  end
       fun mk_snd tm =
           let val ty as Type(@{type_name Product_Type.prod}, [fty,sty]) = type_of tm
-          in  Const ("Product_Type.snd", ty --> sty) $ tm  end
+          in  Const (@{const_name Product_Type.snd}, ty --> sty) $ tm  end
 in
 fun XFILL tych x vstruct =
   let fun traverse p xocc L =
--- a/src/HOL/Tools/TFL/usyntax.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/TFL/usyntax.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -196,7 +196,7 @@
 
 local
 fun mk_uncurry (xt, yt, zt) =
-    Const(@{const_name prod_case}, (xt --> yt --> zt) --> prod_ty xt yt --> zt)
+    Const(@{const_name case_prod}, (xt --> yt --> zt) --> prod_ty xt yt --> zt)
 fun dest_pair(Const(@{const_name Pair},_) $ M $ N) = {fst=M, snd=N}
   | dest_pair _ = raise USYN_ERR "dest_pair" "not a pair"
 fun is_var (Var _) = true | is_var (Free _) = true | is_var _ = false
@@ -276,7 +276,7 @@
   | dest_pair _ = raise USYN_ERR "dest_pair" "not a pair";
 
 
-local  fun ucheck t = (if #Name (dest_const t) = @{const_name prod_case} then t
+local  fun ucheck t = (if #Name (dest_const t) = @{const_name case_prod} then t
                        else raise Match)
 in
 fun dest_pabs used tm =
--- a/src/HOL/Tools/hologic.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/hologic.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -348,21 +348,21 @@
 
 fun mk_fst p =
   let val pT = fastype_of p in
-    Const ("Product_Type.fst", pT --> fst (dest_prodT pT)) $ p
+    Const ("Product_Type.prod.fst", pT --> fst (dest_prodT pT)) $ p
   end;
 
 fun mk_snd p =
   let val pT = fastype_of p in
-    Const ("Product_Type.snd", pT --> snd (dest_prodT pT)) $ p
+    Const ("Product_Type.prod.snd", pT --> snd (dest_prodT pT)) $ p
   end;
 
 fun split_const (A, B, C) =
-  Const ("Product_Type.prod.prod_case", (A --> B --> C) --> mk_prodT (A, B) --> C);
+  Const ("Product_Type.prod.case_prod", (A --> B --> C) --> mk_prodT (A, B) --> C);
 
 fun mk_split t =
   (case Term.fastype_of t of
     T as (Type ("fun", [A, Type ("fun", [B, C])])) =>
-      Const ("Product_Type.prod.prod_case", T --> mk_prodT (A, B) --> C) $ t
+      Const ("Product_Type.prod.case_prod", T --> mk_prodT (A, B) --> C) $ t
   | _ => raise TERM ("mk_split: bad body type", [t]));
 
 (*Maps the type T1 * ... * Tn to [T1, ..., Tn], however nested*)
@@ -478,7 +478,7 @@
 val strip_psplits =
   let
     fun strip [] qs Ts t = (t, rev Ts, qs)
-      | strip (p :: ps) qs Ts (Const ("Product_Type.prod.prod_case", _) $ t) =
+      | strip (p :: ps) qs Ts (Const ("Product_Type.prod.case_prod", _) $ t) =
           strip ((1 :: p) :: (2 :: p) :: ps) (p :: qs) Ts t
       | strip (p :: ps) qs Ts (Abs (s, T, t)) = strip ps qs (T :: Ts) t
       | strip (p :: ps) qs Ts t = strip ps qs
--- a/src/HOL/Tools/set_comprehension_pointfree.ML	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Tools/set_comprehension_pointfree.ML	Wed Feb 12 10:59:25 2014 +0100
@@ -84,7 +84,7 @@
   end;
 
 fun mk_split_abs vs (Bound i) t = let val (x, T) = nth vs i in Abs (x, T, t) end
-  | mk_split_abs vs (Const ("Product_Type.Pair", _) $ u $ v) t =
+  | mk_split_abs vs (Const (@{const_name Product_Type.Pair}, _) $ u $ v) t =
       HOLogic.mk_split (mk_split_abs vs u (mk_split_abs vs v t))
   | mk_split_abs _ t _ = raise TERM ("mk_split_abs: bad term", [t]);
 
@@ -92,7 +92,7 @@
 val strip_psplits =
   let
     fun strip [] qs vs t = (t, rev vs, qs)
-      | strip (p :: ps) qs vs (Const ("Product_Type.prod.prod_case", _) $ t) =
+      | strip (p :: ps) qs vs (Const (@{const_name Product_Type.prod.case_prod}, _) $ t) =
           strip ((1 :: p) :: (2 :: p) :: ps) (p :: qs) vs t
       | strip (_ :: ps) qs vs (Abs (s, T, t)) = strip ps qs ((s, T) :: vs) t
       | strip (_ :: ps) qs vs t = strip ps qs
@@ -305,7 +305,7 @@
 
 (* proof tactic *)
 
-val prod_case_distrib = @{lemma "(prod_case g x) z = prod_case (% x y. (g x y) z) x" by (simp add: prod_case_beta)}
+val case_prod_distrib = @{lemma "(case_prod g x) z = case_prod (% x y. (g x y) z) x" by (simp add: case_prod_beta)}
 
 val vimageI2' = @{lemma "f a \<notin> A ==> a \<notin> f -` A" by simp}
 val vimageE' =
@@ -326,7 +326,7 @@
         @{thm arg_cong2[OF refl, where f="op =", OF prod.cases, THEN iffD2]}
       ORELSE' CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1
         (HOLogic.Trueprop_conv
-          (HOLogic.eq_conv Conv.all_conv (Conv.rewr_conv (mk_meta_eq prod_case_distrib)))))) ctxt)))
+          (HOLogic.eq_conv Conv.all_conv (Conv.rewr_conv (mk_meta_eq case_prod_distrib)))))) ctxt)))
 
 fun elim_image_tac ctxt = etac @{thm imageE}
   THEN' REPEAT_DETERM o CHANGED o
@@ -521,4 +521,3 @@
   end;
 
 end;
-
--- a/src/HOL/Topological_Spaces.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Topological_Spaces.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -1296,9 +1296,9 @@
 proof cases
   let "?P p n" = "p > n \<and> (\<forall>m\<ge>p. s m \<le> s p)"
   assume *: "\<forall>n. \<exists>p. ?P p n"
-  def f \<equiv> "nat_rec (SOME p. ?P p 0) (\<lambda>_ n. SOME p. ?P p n)"
+  def f \<equiv> "rec_nat (SOME p. ?P p 0) (\<lambda>_ n. SOME p. ?P p n)"
   have f_0: "f 0 = (SOME p. ?P p 0)" unfolding f_def by simp
-  have f_Suc: "\<And>i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc ..
+  have f_Suc: "\<And>i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat.recs(2) ..
   have P_0: "?P (f 0) 0" unfolding f_0 using *[rule_format] by (rule someI2_ex) auto
   have P_Suc: "\<And>i. ?P (f (Suc i)) (f i)" unfolding f_Suc using *[rule_format] by (rule someI2_ex) auto
   then have "subseq f" unfolding subseq_Suc_iff by auto
@@ -1318,9 +1318,9 @@
   let "?P p m" = "m < p \<and> s m < s p"
   assume "\<not> (\<forall>n. \<exists>p>n. (\<forall>m\<ge>p. s m \<le> s p))"
   then obtain N where N: "\<And>p. p > N \<Longrightarrow> \<exists>m>p. s p < s m" by (force simp: not_le le_less)
-  def f \<equiv> "nat_rec (SOME p. ?P p (Suc N)) (\<lambda>_ n. SOME p. ?P p n)"
+  def f \<equiv> "rec_nat (SOME p. ?P p (Suc N)) (\<lambda>_ n. SOME p. ?P p n)"
   have f_0: "f 0 = (SOME p. ?P p (Suc N))" unfolding f_def by simp
-  have f_Suc: "\<And>i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc ..
+  have f_Suc: "\<And>i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat.recs(2) ..
   have P_0: "?P (f 0) (Suc N)"
     unfolding f_0 some_eq_ex[of "\<lambda>p. ?P p (Suc N)"] using N[of "Suc N"] by auto
   { fix i have "N < f i \<Longrightarrow> ?P (f (Suc i)) (f i)"
--- a/src/HOL/Transcendental.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Transcendental.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -624,6 +624,7 @@
        (\<lambda>n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))"
       apply (rule ext)
       apply (case_tac "n", simp)
+      apply (rename_tac nat)
       apply (case_tac "nat", simp)
       apply (simp add: r_neq_0)
       done
@@ -2710,7 +2711,8 @@
 apply (subgoal_tac "x < real(LEAST m::nat. x < real m * y) * y")
  prefer 2 apply (erule LeastI)
 apply (case_tac "LEAST m::nat. x < real m * y", simp)
-apply (subgoal_tac "~ x < real nat * y")
+apply (rename_tac m)
+apply (subgoal_tac "~ x < real m * y")
  prefer 2 apply (rule not_less_Least, simp, force)
 done
 
--- a/src/HOL/Transfer.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Transfer.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -358,12 +358,12 @@
   shows "((A ===> B) ===> A ===> B ===> A ===> B) fun_upd fun_upd"
   unfolding fun_upd_def [abs_def] by transfer_prover
 
-lemma nat_case_transfer [transfer_rule]:
-  "(A ===> (op = ===> A) ===> op = ===> A) nat_case nat_case"
+lemma case_nat_transfer [transfer_rule]:
+  "(A ===> (op = ===> A) ===> op = ===> A) case_nat case_nat"
   unfolding fun_rel_def by (simp split: nat.split)
 
-lemma nat_rec_transfer [transfer_rule]:
-  "(A ===> (op = ===> A ===> A) ===> op = ===> A) nat_rec nat_rec"
+lemma rec_nat_transfer [transfer_rule]:
+  "(A ===> (op = ===> A ===> A) ===> op = ===> A) rec_nat rec_nat"
   unfolding fun_rel_def by (clarsimp, rename_tac n, induct_tac n, simp_all)
 
 lemma funpow_transfer [transfer_rule]:
--- a/src/HOL/Transitive_Closure.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Transitive_Closure.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -846,6 +846,7 @@
      \<Longrightarrow> (\<And>y m. n = Suc m \<Longrightarrow> (x, y) \<in> R \<Longrightarrow> (y, z) \<in> R ^^ m \<Longrightarrow> P)
    \<Longrightarrow> P"
   apply (cases n, simp)
+  apply (rename_tac nat)
   apply (cut_tac n=nat and R=R in relpow_Suc_D2', simp, blast)
   done
 
@@ -1297,4 +1298,3 @@
   {* simple transitivity reasoner (predicate version) *}
 
 end
-
--- a/src/HOL/Word/Word.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Word/Word.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -4578,7 +4578,7 @@
 
 definition word_rec :: "'a \<Rightarrow> ('b::len word \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b word \<Rightarrow> 'a"
 where
-  "word_rec forZero forSuc n = nat_rec forZero (forSuc \<circ> of_nat) (unat n)"
+  "word_rec forZero forSuc n = rec_nat forZero (forSuc \<circ> of_nat) (unat n)"
 
 lemma word_rec_0: "word_rec z s 0 = z"
   by (simp add: word_rec_def)
--- a/src/HOL/Word/Word_Miscellaneous.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/Word/Word_Miscellaneous.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -118,10 +118,6 @@
 
 lemmas seqr = eq_reflection [where x = "size w"] for w (* FIXME: delete *)
 
-(* TODO: move name bindings to List.thy *)
-lemmas tl_Nil = tl.simps (1)
-lemmas tl_Cons = tl.simps (2)
-
 lemma the_elemI: "y = {x} ==> the_elem y = x" 
   by simp
 
--- a/src/HOL/ex/Primrec.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/ex/Primrec.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -191,7 +191,7 @@
   "PREC f g l =
     (case l of
       [] => 0
-    | x # l' => nat_rec (f l') (\<lambda>y r. g (r # y # l')) x)"
+    | x # l' => rec_nat (f l') (\<lambda>y r. g (r # y # l')) x)"
   -- {* Note that @{term g} is applied first to @{term "PREC f g y"} and then to @{term y}! *}
 
 inductive PRIMREC :: "(nat list => nat) => bool" where
--- a/src/HOL/ex/Refute_Examples.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/ex/Refute_Examples.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -525,14 +525,6 @@
 refute [expect = genuine]
 oops
 
-lemma "unit_rec u x = u"
-refute [expect = none]
-by simp
-
-lemma "P (unit_rec u x)"
-refute [expect = genuine]
-oops
-
 lemma "P (case x of () \<Rightarrow> u)"
 refute [expect = genuine]
 oops
@@ -555,15 +547,15 @@
 refute [expect = genuine]
 oops
 
-lemma "option_rec n s None = n"
+lemma "rec_option n s None = n"
 refute [expect = none]
 by simp
 
-lemma "option_rec n s (Some x) = s x"
+lemma "rec_option n s (Some x) = s x"
 refute [maxsize = 4, expect = none]
 by simp
 
-lemma "P (option_rec n s x)"
+lemma "P (rec_option n s x)"
 refute [expect = genuine]
 oops
 
@@ -597,14 +589,6 @@
 refute [expect = genuine]
 oops
 
-lemma "prod_rec p (a, b) = p a b"
-refute [maxsize = 2, expect = none]
-by simp
-
-lemma "P (prod_rec p x)"
-refute [expect = genuine]
-oops
-
 lemma "P (case x of Pair a b \<Rightarrow> p a b)"
 refute [expect = genuine]
 oops
@@ -631,18 +615,6 @@
 refute [expect = genuine]
 oops
 
-lemma "sum_rec l r (Inl x) = l x"
-refute [maxsize = 3, expect = none]
-by simp
-
-lemma "sum_rec l r (Inr x) = r x"
-refute [maxsize = 3, expect = none]
-by simp
-
-lemma "P (sum_rec l r x)"
-refute [expect = genuine]
-oops
-
 lemma "P (case x of Inl a \<Rightarrow> l a | Inr b \<Rightarrow> r b)"
 refute [expect = genuine]
 oops
@@ -667,15 +639,15 @@
 refute [expect = genuine]
 oops
 
-lemma "T1_rec a b A = a"
+lemma "rec_T1 a b A = a"
 refute [expect = none]
 by simp
 
-lemma "T1_rec a b B = b"
+lemma "rec_T1 a b B = b"
 refute [expect = none]
 by simp
 
-lemma "P (T1_rec a b x)"
+lemma "P (rec_T1 a b x)"
 refute [expect = genuine]
 oops
 
@@ -697,15 +669,15 @@
 refute [expect = genuine]
 oops
 
-lemma "T2_rec c d (C x) = c x"
+lemma "rec_T2 c d (C x) = c x"
 refute [maxsize = 4, expect = none]
 by simp
 
-lemma "T2_rec c d (D x) = d x"
+lemma "rec_T2 c d (D x) = d x"
 refute [maxsize = 4, expect = none]
 by simp
 
-lemma "P (T2_rec c d x)"
+lemma "P (rec_T2 c d x)"
 refute [expect = genuine]
 oops
 
@@ -727,11 +699,11 @@
 refute [expect = genuine]
 oops
 
-lemma "T3_rec e (E x) = e x"
+lemma "rec_T3 e (E x) = e x"
 refute [maxsize = 2, expect = none]
 by simp
 
-lemma "P (T3_rec e x)"
+lemma "P (rec_T3 e x)"
 refute [expect = genuine]
 oops
 
@@ -762,15 +734,15 @@
       model will be found *}
 oops
 
-lemma "nat_rec zero suc 0 = zero"
+lemma "rec_nat zero suc 0 = zero"
 refute [expect = none]
 by simp
 
-lemma "nat_rec zero suc (Suc x) = suc x (nat_rec zero suc x)"
+lemma "rec_nat zero suc (Suc x) = suc x (rec_nat zero suc x)"
 refute [maxsize = 2, expect = none]
 by simp
 
-lemma "P (nat_rec zero suc x)"
+lemma "P (rec_nat zero suc x)"
 refute [expect = potential]
 oops
 
@@ -792,15 +764,15 @@
 refute [expect = potential]
 oops
 
-lemma "list_rec nil cons [] = nil"
+lemma "rec_list nil cons [] = nil"
 refute [maxsize = 3, expect = none]
 by simp
 
-lemma "list_rec nil cons (x#xs) = cons x xs (list_rec nil cons xs)"
+lemma "rec_list nil cons (x#xs) = cons x xs (rec_list nil cons xs)"
 refute [maxsize = 2, expect = none]
 by simp
 
-lemma "P (list_rec nil cons xs)"
+lemma "P (rec_list nil cons xs)"
 refute [expect = potential]
 oops
 
@@ -830,19 +802,19 @@
 refute [expect = potential]
 oops
 
-lemma "BitList_rec nil bit0 bit1 BitListNil = nil"
+lemma "rec_BitList nil bit0 bit1 BitListNil = nil"
 refute [maxsize = 4, expect = none]
 by simp
 
-lemma "BitList_rec nil bit0 bit1 (Bit0 xs) = bit0 xs (BitList_rec nil bit0 bit1 xs)"
+lemma "rec_BitList nil bit0 bit1 (Bit0 xs) = bit0 xs (rec_BitList nil bit0 bit1 xs)"
 refute [maxsize = 2, expect = none]
 by simp
 
-lemma "BitList_rec nil bit0 bit1 (Bit1 xs) = bit1 xs (BitList_rec nil bit0 bit1 xs)"
+lemma "rec_BitList nil bit0 bit1 (Bit1 xs) = bit1 xs (rec_BitList nil bit0 bit1 xs)"
 refute [maxsize = 2, expect = none]
 by simp
 
-lemma "P (BitList_rec nil bit0 bit1 x)"
+lemma "P (rec_BitList nil bit0 bit1 x)"
 refute [expect = potential]
 oops
 
@@ -860,7 +832,7 @@
 refute [expect = potential]
 oops
 
-lemma "BinTree_rec l n (Leaf x) = l x"
+lemma "rec_BinTree l n (Leaf x) = l x"
   refute [maxsize = 1, expect = none]
   (* The "maxsize = 1" tests are a bit pointless: for some formulae
      below, refute will find no countermodel simply because this
@@ -868,11 +840,11 @@
      larger size already takes too long. *)
 by simp
 
-lemma "BinTree_rec l n (Node x y) = n x y (BinTree_rec l n x) (BinTree_rec l n y)"
+lemma "rec_BinTree l n (Node x y) = n x y (rec_BinTree l n x) (rec_BinTree l n y)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "P (BinTree_rec l n x)"
+lemma "P (rec_BinTree l n x)"
 refute [expect = potential]
 oops
 
@@ -905,15 +877,15 @@
 refute [expect = potential]
 oops
 
-lemma "aexp_bexp_rec_1 number ite equal (Number x) = number x"
+lemma "rec_aexp_bexp_1 number ite equal (Number x) = number x"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "aexp_bexp_rec_1 number ite equal (ITE x y z) = ite x y z (aexp_bexp_rec_2 number ite equal x) (aexp_bexp_rec_1 number ite equal y) (aexp_bexp_rec_1 number ite equal z)"
+lemma "rec_aexp_bexp_1 number ite equal (ITE x y z) = ite x y z (rec_aexp_bexp_2 number ite equal x) (rec_aexp_bexp_1 number ite equal y) (rec_aexp_bexp_1 number ite equal z)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "P (aexp_bexp_rec_1 number ite equal x)"
+lemma "P (rec_aexp_bexp_1 number ite equal x)"
 refute [expect = potential]
 oops
 
@@ -921,11 +893,11 @@
 refute [expect = potential]
 oops
 
-lemma "aexp_bexp_rec_2 number ite equal (Equal x y) = equal x y (aexp_bexp_rec_1 number ite equal x) (aexp_bexp_rec_1 number ite equal y)"
+lemma "rec_aexp_bexp_2 number ite equal (Equal x y) = equal x y (rec_aexp_bexp_1 number ite equal x) (rec_aexp_bexp_1 number ite equal y)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "P (aexp_bexp_rec_2 number ite equal x)"
+lemma "P (rec_aexp_bexp_2 number ite equal x)"
 refute [expect = potential]
 oops
 
@@ -980,35 +952,35 @@
 refute [expect = potential]
 oops
 
-lemma "X_Y_rec_1 a b c d e f A = a"
+lemma "rec_X_Y_1 a b c d e f A = a"
 refute [maxsize = 3, expect = none]
 by simp
 
-lemma "X_Y_rec_1 a b c d e f (B x) = b x (X_Y_rec_1 a b c d e f x)"
+lemma "rec_X_Y_1 a b c d e f (B x) = b x (rec_X_Y_1 a b c d e f x)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "X_Y_rec_1 a b c d e f (C y) = c y (X_Y_rec_2 a b c d e f y)"
+lemma "rec_X_Y_1 a b c d e f (C y) = c y (rec_X_Y_2 a b c d e f y)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "X_Y_rec_2 a b c d e f (D x) = d x (X_Y_rec_1 a b c d e f x)"
+lemma "rec_X_Y_2 a b c d e f (D x) = d x (rec_X_Y_1 a b c d e f x)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "X_Y_rec_2 a b c d e f (E y) = e y (X_Y_rec_2 a b c d e f y)"
+lemma "rec_X_Y_2 a b c d e f (E y) = e y (rec_X_Y_2 a b c d e f y)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "X_Y_rec_2 a b c d e f F = f"
+lemma "rec_X_Y_2 a b c d e f F = f"
 refute [maxsize = 3, expect = none]
 by simp
 
-lemma "P (X_Y_rec_1 a b c d e f x)"
+lemma "P (rec_X_Y_1 a b c d e f x)"
 refute [expect = potential]
 oops
 
-lemma "P (X_Y_rec_2 a b c d e f y)"
+lemma "P (rec_X_Y_2 a b c d e f y)"
 refute [expect = potential]
 oops
 
@@ -1030,39 +1002,39 @@
 refute [expect = potential]
 oops
 
-lemma "XOpt_rec_1 cx dx n1 s1 n2 s2 (CX x) = cx x (XOpt_rec_2 cx dx n1 s1 n2 s2 x)"
+lemma "rec_XOpt_1 cx dx n1 s1 n2 s2 (CX x) = cx x (rec_XOpt_2 cx dx n1 s1 n2 s2 x)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "XOpt_rec_1 cx dx n1 s1 n2 s2 (DX x) = dx x (\<lambda>b. XOpt_rec_3 cx dx n1 s1 n2 s2 (x b))"
+lemma "rec_XOpt_1 cx dx n1 s1 n2 s2 (DX x) = dx x (\<lambda>b. rec_XOpt_3 cx dx n1 s1 n2 s2 (x b))"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "XOpt_rec_2 cx dx n1 s1 n2 s2 None = n1"
+lemma "rec_XOpt_2 cx dx n1 s1 n2 s2 None = n1"
 refute [maxsize = 2, expect = none]
 by simp
 
-lemma "XOpt_rec_2 cx dx n1 s1 n2 s2 (Some x) = s1 x (XOpt_rec_1 cx dx n1 s1 n2 s2 x)"
+lemma "rec_XOpt_2 cx dx n1 s1 n2 s2 (Some x) = s1 x (rec_XOpt_1 cx dx n1 s1 n2 s2 x)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "XOpt_rec_3 cx dx n1 s1 n2 s2 None = n2"
+lemma "rec_XOpt_3 cx dx n1 s1 n2 s2 None = n2"
 refute [maxsize = 2, expect = none]
 by simp
 
-lemma "XOpt_rec_3 cx dx n1 s1 n2 s2 (Some x) = s2 x (XOpt_rec_1 cx dx n1 s1 n2 s2 x)"
+lemma "rec_XOpt_3 cx dx n1 s1 n2 s2 (Some x) = s2 x (rec_XOpt_1 cx dx n1 s1 n2 s2 x)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "P (XOpt_rec_1 cx dx n1 s1 n2 s2 x)"
+lemma "P (rec_XOpt_1 cx dx n1 s1 n2 s2 x)"
 refute [expect = potential]
 oops
 
-lemma "P (XOpt_rec_2 cx dx n1 s1 n2 s2 x)"
+lemma "P (rec_XOpt_2 cx dx n1 s1 n2 s2 x)"
 refute [expect = potential]
 oops
 
-lemma "P (XOpt_rec_3 cx dx n1 s1 n2 s2 x)"
+lemma "P (rec_XOpt_3 cx dx n1 s1 n2 s2 x)"
 refute [expect = potential]
 oops
 
@@ -1080,23 +1052,23 @@
 refute [expect = potential]
 oops
 
-lemma "YOpt_rec_1 cy n s (CY x) = cy x (YOpt_rec_2 cy n s x)"
+lemma "rec_YOpt_1 cy n s (CY x) = cy x (rec_YOpt_2 cy n s x)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "YOpt_rec_2 cy n s None = n"
+lemma "rec_YOpt_2 cy n s None = n"
 refute [maxsize = 2, expect = none]
 by simp
 
-lemma "YOpt_rec_2 cy n s (Some x) = s x (\<lambda>a. YOpt_rec_1 cy n s (x a))"
+lemma "rec_YOpt_2 cy n s (Some x) = s x (\<lambda>a. rec_YOpt_1 cy n s (x a))"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "P (YOpt_rec_1 cy n s x)"
+lemma "P (rec_YOpt_1 cy n s x)"
 refute [expect = potential]
 oops
 
-lemma "P (YOpt_rec_2 cy n s x)"
+lemma "P (rec_YOpt_2 cy n s x)"
 refute [expect = potential]
 oops
 
@@ -1114,23 +1086,23 @@
 refute [expect = potential]
 oops
 
-lemma "Trie_rec_1 tr nil cons (TR x) = tr x (Trie_rec_2 tr nil cons x)"
+lemma "rec_Trie_1 tr nil cons (TR x) = tr x (rec_Trie_2 tr nil cons x)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "Trie_rec_2 tr nil cons [] = nil"
+lemma "rec_Trie_2 tr nil cons [] = nil"
 refute [maxsize = 3, expect = none]
 by simp
 
-lemma "Trie_rec_2 tr nil cons (x#xs) = cons x xs (Trie_rec_1 tr nil cons x) (Trie_rec_2 tr nil cons xs)"
+lemma "rec_Trie_2 tr nil cons (x#xs) = cons x xs (rec_Trie_1 tr nil cons x) (rec_Trie_2 tr nil cons xs)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "P (Trie_rec_1 tr nil cons x)"
+lemma "P (rec_Trie_1 tr nil cons x)"
 refute [expect = potential]
 oops
 
-lemma "P (Trie_rec_2 tr nil cons x)"
+lemma "P (rec_Trie_2 tr nil cons x)"
 refute [expect = potential]
 oops
 
@@ -1148,15 +1120,15 @@
 refute [expect = potential]
 oops
 
-lemma "InfTree_rec leaf node Leaf = leaf"
+lemma "rec_InfTree leaf node Leaf = leaf"
 refute [maxsize = 2, expect = none]
 by simp
 
-lemma "InfTree_rec leaf node (Node x) = node x (\<lambda>n. InfTree_rec leaf node (x n))"
+lemma "rec_InfTree leaf node (Node x) = node x (\<lambda>n. rec_InfTree leaf node (x n))"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "P (InfTree_rec leaf node x)"
+lemma "P (rec_InfTree leaf node x)"
 refute [expect = potential]
 oops
 
@@ -1174,19 +1146,19 @@
 refute [expect = potential]
 oops
 
-lemma "lambda_rec var app lam (Var x) = var x"
+lemma "rec_lambda var app lam (Var x) = var x"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "lambda_rec var app lam (App x y) = app x y (lambda_rec var app lam x) (lambda_rec var app lam y)"
+lemma "rec_lambda var app lam (App x y) = app x y (rec_lambda var app lam x) (rec_lambda var app lam y)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "lambda_rec var app lam (Lam x) = lam x (\<lambda>a. lambda_rec var app lam (x a))"
+lemma "rec_lambda var app lam (Lam x) = lam x (\<lambda>a. rec_lambda var app lam (x a))"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "P (lambda_rec v a l x)"
+lemma "P (rec_lambda v a l x)"
 refute [expect = potential]
 oops
 
@@ -1207,35 +1179,35 @@
 refute [expect = potential]
 oops
 
-lemma "U_rec_1 e c d nil cons (E x) = e x (U_rec_2 e c d nil cons x)"
+lemma "rec_U_1 e c d nil cons (E x) = e x (rec_U_2 e c d nil cons x)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "U_rec_2 e c d nil cons (C x) = c x"
+lemma "rec_U_2 e c d nil cons (C x) = c x"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "U_rec_2 e c d nil cons (D x) = d x (U_rec_3 e c d nil cons x)"
+lemma "rec_U_2 e c d nil cons (D x) = d x (rec_U_3 e c d nil cons x)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "U_rec_3 e c d nil cons [] = nil"
+lemma "rec_U_3 e c d nil cons [] = nil"
 refute [maxsize = 2, expect = none]
 by simp
 
-lemma "U_rec_3 e c d nil cons (x#xs) = cons x xs (U_rec_1 e c d nil cons x) (U_rec_3 e c d nil cons xs)"
+lemma "rec_U_3 e c d nil cons (x#xs) = cons x xs (rec_U_1 e c d nil cons x) (rec_U_3 e c d nil cons xs)"
 refute [maxsize = 1, expect = none]
 by simp
 
-lemma "P (U_rec_1 e c d nil cons x)"
+lemma "P (rec_U_1 e c d nil cons x)"
 refute [expect = potential]
 oops
 
-lemma "P (U_rec_2 e c d nil cons x)"
+lemma "P (rec_U_2 e c d nil cons x)"
 refute [expect = potential]
 oops
 
-lemma "P (U_rec_3 e c d nil cons x)"
+lemma "P (rec_U_3 e c d nil cons x)"
 refute [expect = potential]
 oops
 
--- a/src/HOL/ex/Tree23.thy	Wed Feb 12 09:06:04 2014 +0100
+++ b/src/HOL/ex/Tree23.thy	Wed Feb 12 10:59:25 2014 +0100
@@ -330,12 +330,12 @@
 "dfull n (Some (p, (False, t'))) \<longleftrightarrow> full (Suc n) t'"
 
 lemmas dfull_case_intros =
-  ord.exhaust [where y=y and P="dfull a (ord_case b c d y)"]
-  option.exhaust [where y=y and P="dfull a (option_case b c y)"]
-  prod.exhaust [where y=y and P="dfull a (prod_case b y)"]
-  bool.exhaust [where y=y and P="dfull a (bool_case b c y)"]
-  tree23.exhaust [where y=y and P="dfull a (Some (b, tree23_case c d e y))"]
-  tree23.exhaust [where y=y and P="full a (tree23_case b c d y)"]
+  ord.exhaust [of y "dfull a (case_ord b c d y)"]
+  option.exhaust [of y "dfull a (case_option b c y)"]
+  prod.exhaust [of y "dfull a (case_prod b y)"]
+  bool.exhaust [of y "dfull a (case_bool b c y)"]
+  tree23.exhaust [of y "dfull a (Some (b, case_tree23 c d e y))"]
+  tree23.exhaust [of y "full a (case_tree23 b c d y)"]
   for a b c d e y
 
 lemma dfull_del: "full (Suc n) t \<Longrightarrow> dfull n (del k t)"
@@ -396,7 +396,7 @@
 apply (case_tac n, simp, simp)
 apply (frule dfull_del [where k="Some k"])
 apply (cases "del (Some k) t", force)
-apply (case_tac "a", rename_tac p b t', case_tac "b", auto)
+apply (rename_tac a, case_tac "a", rename_tac b t', case_tac "b", auto)
 done
 
 text{* This is a little test harness and should be commented out once the