merged, using src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML and src/HOL/Tools/Sledgehammer/sledgehammer_run.ML from 347c3b0cab44;
authorwenzelm
Mon, 11 Nov 2013 17:44:21 +0100
changeset 54384 50199af40c27
parent 54383 9d3c7a04a65e (current diff)
parent 54298 347c3b0cab44 (diff)
child 54385 27246f8b2dac
merged, using src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML and src/HOL/Tools/Sledgehammer/sledgehammer_run.ML from 347c3b0cab44;
CONTRIBUTORS
NEWS
src/Doc/manual.bib
src/HOL/BNF/Tools/bnf_fp_rec_sugar.ML
src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML
src/HOL/Library/Abstract_Rat.thy
src/HOL/Library/Glbs.thy
src/HOL/Library/Univ_Poly.thy
src/HOL/Lubs.thy
--- a/CONTRIBUTORS	Mon Nov 11 17:34:44 2013 +0100
+++ b/CONTRIBUTORS	Mon Nov 11 17:44:21 2013 +0100
@@ -3,6 +3,10 @@
 who is listed as an author in one of the source files of this Isabelle
 distribution.
 
+Contributions to this Isabelle version
+--------------------------------------
+
+
 Contributions to Isabelle2013-1
 -------------------------------
 
--- a/NEWS	Mon Nov 11 17:34:44 2013 +0100
+++ b/NEWS	Mon Nov 11 17:44:21 2013 +0100
@@ -1,6 +1,67 @@
 Isabelle NEWS -- history user-relevant changes
 ==============================================
 
+New in this Isabelle version
+----------------------------
+
+*** HOL ***
+
+* Qualified constant names Wellfounded.acc, Wellfounded.accp.
+INCOMPATIBILITY.
+
+* Fact generalization and consolidation:
+    neq_one_mod_two, mod_2_not_eq_zero_eq_one_int ~> not_mod_2_eq_0_eq_1
+INCOMPATIBILITY.
+
+* Purely algebraic definition of even.  Fact generalization and consolidation:
+    nat_even_iff_2_dvd, int_even_iff_2_dvd ~> even_iff_2_dvd
+    even_zero_(nat|int) ~> even_zero
+INCOMPATIBILITY.
+
+* Elimination of fact duplicates:
+    equals_zero_I ~> minus_unique
+    diff_eq_0_iff_eq ~> right_minus_eq
+INCOMPATIBILITY.
+
+* Fact name consolidation:
+    diff_def, diff_minus, ab_diff_minus ~> diff_conv_add_uminus
+    minus_le_self_iff ~> neg_less_eq_nonneg
+    le_minus_self_iff ~> less_eq_neg_nonpos
+    neg_less_nonneg ~> neg_less_pos
+    less_minus_self_iff ~> less_neg_neg [simp]
+INCOMPATIBILITY.
+
+* More simplification rules on unary and binary minus:
+add_diff_cancel, add_diff_cancel_left, add_le_same_cancel1,
+add_le_same_cancel2, add_less_same_cancel1, add_less_same_cancel2,
+add_minus_cancel, diff_add_cancel, le_add_same_cancel1,
+le_add_same_cancel2, less_add_same_cancel1, less_add_same_cancel2,
+minus_add_cancel, uminus_add_conv_diff.  These correspondingly
+have been taken away from fact collections algebra_simps and
+field_simps.  INCOMPATIBILITY.
+
+To restore proofs, the following patterns are helpful:
+
+a) Arbitrary failing proof not involving "diff_def":
+Consider simplification with algebra_simps or field_simps.
+
+b) Lifting rules from addition to subtraction:
+Try with "using <rule for addition> of [… "- _" …]" by simp".
+
+c) Simplification with "diff_def": just drop "diff_def".
+Consider simplification with algebra_simps or field_simps;
+or the brute way with
+"simp add: diff_conv_add_uminus del: add_uminus_conv_diff".
+
+* SUP and INF generalized to conditionally_complete_lattice
+
+* Theory Lubs moved HOL image to HOL-Library. It is replaced by
+Conditionally_Complete_Lattices.   INCOMPATIBILITY.
+
+* Introduce bdd_above and bdd_below in Conditionally_Complete_Lattices, use them
+instead of explicitly stating boundedness of sets.
+
+
 New in Isabelle2013-1 (November 2013)
 -------------------------------------
 
--- a/src/Doc/Datatypes/Datatypes.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/Doc/Datatypes/Datatypes.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -9,21 +9,8 @@
 
 theory Datatypes
 imports Setup
-keywords
-  "primcorec_notyet" :: thy_decl
 begin
 
-(*<*)
-(* FIXME: Temporary setup until "primcorec" and "primcorecursive" are fully implemented. *)
-ML_command {*
-fun add_dummy_cmd _ _ lthy = lthy;
-
-val _ = Outer_Syntax.local_theory @{command_spec "primcorec_notyet"} ""
-  (Parse.fixes -- Parse_Spec.where_alt_specs >> uncurry add_dummy_cmd);
-*}
-(*>*)
-
-
 section {* Introduction
   \label{sec:introduction} *}
 
@@ -54,17 +41,19 @@
 
 text {*
 \noindent
-The package also provides some convenience, notably automatically generated
-discriminators and selectors.
-
-In addition to plain inductive datatypes, the new package supports coinductive
-datatypes, or \emph{codatatypes}, which may have infinite values. For example,
-the following command introduces the type of lazy lists, which comprises both
-finite and infinite values:
+Furthermore, the package provides a lot of convenience, including automatically
+generated discriminators, selectors, and relators as well as a wealth of
+properties about them.
+
+In addition to inductive datatypes, the new package supports coinductive
+datatypes, or \emph{codatatypes}, which allow infinite values. For example, the
+following command introduces the type of lazy lists, which comprises both finite
+and infinite values:
 *}
 
 (*<*)
     locale early
+    locale late
 (*>*)
     codatatype (*<*)(in early) (*>*)'a llist = LNil | LCons 'a "'a llist"
 
@@ -80,10 +69,10 @@
     codatatype (*<*)(in early) (*>*)'a tree\<^sub>i\<^sub>i = Node\<^sub>i\<^sub>i 'a "'a tree\<^sub>i\<^sub>i llist"
 
 text {*
-The first two tree types allow only finite branches, whereas the last two allow
-branches of infinite length. Orthogonally, the nodes in the first and third
-types have finite branching, whereas those of the second and fourth may have
-infinitely many direct subtrees.
+The first two tree types allow only paths of finite length, whereas the last two
+allow infinite paths. Orthogonally, the nodes in the first and third types have
+finitely many direct subtrees, whereas those of the second and fourth may have
+infinite branching.
 
 To use the package, it is necessary to import the @{theory BNF} theory, which
 can be precompiled into the \texttt{HOL-BNF} image. The following commands show
@@ -152,15 +141,15 @@
 
 
 \newbox\boxA
-\setbox\boxA=\hbox{\texttt{nospam}}
-
-\newcommand\authoremaili{\texttt{blan{\color{white}nospam}\kern-\wd\boxA{}chette@\allowbreak
+\setbox\boxA=\hbox{\texttt{NOSPAM}}
+
+\newcommand\authoremaili{\texttt{blan{\color{white}NOSPAM}\kern-\wd\boxA{}chette@\allowbreak
 in.\allowbreak tum.\allowbreak de}}
-\newcommand\authoremailii{\texttt{lore{\color{white}nospam}\kern-\wd\boxA{}nz.panny@\allowbreak
+\newcommand\authoremailii{\texttt{lore{\color{white}NOSPAM}\kern-\wd\boxA{}nz.panny@\allowbreak
 \allowbreak tum.\allowbreak de}}
-\newcommand\authoremailiii{\texttt{pope{\color{white}nospam}\kern-\wd\boxA{}scua@\allowbreak
+\newcommand\authoremailiii{\texttt{pope{\color{white}NOSPAM}\kern-\wd\boxA{}scua@\allowbreak
 in.\allowbreak tum.\allowbreak de}}
-\newcommand\authoremailiv{\texttt{tray{\color{white}nospam}\kern-\wd\boxA{}tel@\allowbreak
+\newcommand\authoremailiv{\texttt{tray{\color{white}NOSPAM}\kern-\wd\boxA{}tel@\allowbreak
 in.\allowbreak tum.\allowbreak de}}
 
 The commands @{command datatype_new} and @{command primrec_new} are expected to
@@ -171,13 +160,6 @@
 Comments and bug reports concerning either the tool or this tutorial should be
 directed to the authors at \authoremaili, \authoremailii, \authoremailiii,
 and \authoremailiv.
-
-\begin{framed}
-\noindent
-\textbf{Warning:}\enskip This tutorial and the package it describes are under
-construction. Please forgive their appearance. Should you have suggestions
-or comments regarding either, please let the authors know.
-\end{framed}
 *}
 
 
@@ -195,7 +177,7 @@
 text {*
 Datatypes are illustrated through concrete examples featuring different flavors
 of recursion. More examples can be found in the directory
-\verb|~~/src/HOL/BNF/Examples|.
+\verb|~~/src/HOL/|\allowbreak\verb|BNF/Examples|.
 *}
 
 subsubsection {* Nonrecursive Types
@@ -260,7 +242,8 @@
 
 text {*
 \noindent
-Lists were shown in the introduction. Terminated lists are a variant:
+Lists were shown in the introduction. Terminated lists are a variant that
+stores a value of type @{typ 'b} at the very end:
 *}
 
     datatype_new (*<*)(in early) (*>*)('a, 'b) tlist = TNil 'b | TCons 'a "('a, 'b) tlist"
@@ -310,7 +293,7 @@
 Not all nestings are admissible. For example, this command will fail:
 *}
 
-    datatype_new 'a wrong = Wrong (*<*)'a
+    datatype_new 'a wrong = W1 | W2 (*<*)'a
     typ (*>*)"'a wrong \<Rightarrow> 'a"
 
 text {*
@@ -321,7 +304,7 @@
 *}
 
     datatype_new ('a, 'b) fn = Fn "'a \<Rightarrow> 'b"
-    datatype_new 'a also_wrong = Also_Wrong (*<*)'a
+    datatype_new 'a also_wrong = W1 | W2 (*<*)'a
     typ (*>*)"('a also_wrong, 'a) fn"
 
 text {*
@@ -344,20 +327,30 @@
 datatype_new} and @{command codatatype} commands.
 Section~\ref{sec:registering-bounded-natural-functors} explains how to register
 arbitrary type constructors as BNFs.
+
+Here is another example that fails:
 *}
 
-
-subsubsection {* Custom Names and Syntaxes
-  \label{sssec:datatype-custom-names-and-syntaxes} *}
+    datatype_new 'a pow_list = PNil 'a (*<*)'a
+    datatype_new 'a pow_list' = PNil' 'a (*>*)| PCons "('a * 'a) pow_list"
+
+text {*
+\noindent
+This one features a different flavor of nesting, where the recursive call in the
+type specification occurs around (rather than inside) another type constructor.
+*}
+
+subsubsection {* Auxiliary Constants and Properties
+  \label{sssec:datatype-auxiliary-constants-and-properties} *}
 
 text {*
 The @{command datatype_new} command introduces various constants in addition to
 the constructors. With each datatype are associated set functions, a map
 function, a relator, discriminators, and selectors, all of which can be given
-custom names. In the example below, the traditional names
-@{text set}, @{text map}, @{text list_all2}, @{text null}, @{text hd}, and
-@{text tl} override the default names @{text list_set}, @{text list_map}, @{text
-list_rel}, @{text is_Nil}, @{text un_Cons1}, and @{text un_Cons2}:
+custom names. In the example below, the familiar names @{text null}, @{text hd},
+@{text tl}, @{text set}, @{text map}, and @{text list_all2}, override the
+default names @{text is_Nil}, @{text un_Cons1}, @{text un_Cons2},
+@{text list_set}, @{text list_map}, and @{text list_rel}:
 *}
 
 (*<*)
@@ -380,14 +373,34 @@
 
 text {*
 \noindent
-The command introduces a discriminator @{const null} and a pair of selectors
-@{const hd} and @{const tl} characterized as follows:
+
+\begin{tabular}{@ {}ll@ {}}
+Constructors: &
+  @{text "Nil \<Colon> 'a list"} \\
+&
+  @{text "Cons \<Colon> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list"} \\
+Discriminator: &
+  @{text "null \<Colon> 'a list \<Rightarrow> bool"} \\
+Selectors: &
+  @{text "hd \<Colon> 'a list \<Rightarrow> 'a"} \\
+&
+  @{text "tl \<Colon> 'a list \<Rightarrow> 'a list"} \\
+Set function: &
+  @{text "set \<Colon> 'a list \<Rightarrow> 'a set"} \\
+Map function: &
+  @{text "map \<Colon> ('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list"} \\
+Relator: &
+  @{text "list_all2 \<Colon> ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> bool"}
+\end{tabular}
+
+The discriminator @{const null} and the selectors @{const hd} and @{const tl}
+are characterized as follows:
 %
 \[@{thm list.collapse(1)[of xs, no_vars]}
   \qquad @{thm list.collapse(2)[of xs, no_vars]}\]
 %
-For two-constructor datatypes, a single discriminator constant suffices. The
-discriminator associated with @{const Cons} is simply
+For two-constructor datatypes, a single discriminator constant is sufficient.
+The discriminator associated with @{const Cons} is simply
 @{term "\<lambda>xs. \<not> null xs"}.
 
 The @{text defaults} clause following the @{const Nil} constructor specifies a
@@ -589,6 +602,10 @@
 or the function type. In principle, it should be possible to support old-style
 datatypes as well, but the command does not support this yet (and there is
 currently no way to register old-style datatypes as new-style datatypes).
+
+\item The recursor produced for types that recurse through functions has a
+different signature than with the old package. This makes it impossible to use
+the old \keyw{primrec} command.
 \end{itemize}
 
 An alternative to @{command datatype_new_compat} is to use the old package's
@@ -636,7 +653,7 @@
 \noindent
 The case combinator, discriminators, and selectors are collectively called
 \emph{destructors}. The prefix ``@{text "t."}'' is an optional component of the
-name and is normally hidden. 
+name and is normally hidden.
 *}
 
 
@@ -798,6 +815,10 @@
 
 \end{description}
 \end{indentblock}
+
+\noindent
+In addition, equational versions of @{text t.disc} are registered with the @{text "[code]"}
+attribute.
 *}
 
 
@@ -818,16 +839,20 @@
 @{thm list.map(1)[no_vars]} \\
 @{thm list.map(2)[no_vars]}
 
-\item[@{text "t."}\hthm{rel\_inject} @{text "[simp, code]"}\rm:] ~ \\
+\item[@{text "t."}\hthm{rel\_inject} @{text "[simp]"}\rm:] ~ \\
 @{thm list.rel_inject(1)[no_vars]} \\
 @{thm list.rel_inject(2)[no_vars]}
 
-\item[@{text "t."}\hthm{rel\_distinct} @{text "[simp, code]"}\rm:] ~ \\
+\item[@{text "t."}\hthm{rel\_distinct} @{text "[simp]"}\rm:] ~ \\
 @{thm list.rel_distinct(1)[no_vars]} \\
 @{thm list.rel_distinct(2)[no_vars]}
 
 \end{description}
 \end{indentblock}
+
+\noindent
+In addition, equational versions of @{text t.rel_inject} and @{text
+rel_distinct} are registered with the @{text "[code]"} attribute.
 *}
 
 
@@ -890,17 +915,18 @@
 to register new-style datatypes as old-style datatypes.
 
 \item \emph{The recursor @{text "t_rec"} has a different signature for nested
-recursive datatypes.} In the old package, nested recursion was internally
-reduced to mutual recursion. This reduction was visible in the type of the
-recursor, used by \keyw{primrec}. In the new package, nested recursion is
-handled in a more modular fashion. The old-style recursor can be generated on
-demand using @{command primrec_new}, as explained in
+recursive datatypes.} In the old package, nested recursion through non-functions
+was internally reduced to mutual recursion. This reduction was visible in the
+type of the recursor, used by \keyw{primrec}. Recursion through functions was
+handled specially. In the new package, nested recursion (for functions and
+non-functions) is handled in a more modular fashion. The old-style recursor can
+be generated on demand using @{command primrec_new}, as explained in
 Section~\ref{sssec:primrec-nested-as-mutual-recursion}, if the recursion is via
 new-style datatypes.
 
-\item \emph{Accordingly, the induction principle is different for nested
-recursive datatypes.} Again, the old-style induction principle can be generated
-on demand using @{command primrec_new}, as explained in
+\item \emph{Accordingly, the induction rule is different for nested recursive
+datatypes.} Again, the old-style induction rule can be generated on demand using
+@{command primrec_new}, as explained in
 Section~\ref{sssec:primrec-nested-as-mutual-recursion}, if the recursion is via
 new-style datatypes.
 
@@ -940,9 +966,9 @@
   \label{sec:defining-recursive-functions} *}
 
 text {*
-Recursive functions over datatypes can be specified using @{command
-primrec_new}, which supports primitive recursion, or using the more general
-\keyw{fun} and \keyw{function} commands. Here, the focus is on @{command
+Recursive functions over datatypes can be specified using the @{command
+primrec_new} command, which supports primitive recursion, or using the more
+general \keyw{fun} and \keyw{function} commands. Here, the focus is on @{command
 primrec_new}; the other two commands are described in a separate tutorial
 \cite{isabelle-function}.
 
@@ -1026,9 +1052,10 @@
 
 text {*
 \noindent
-The next example is not primitive recursive, but it can be defined easily using
-\keyw{fun}. The @{command datatype_new_compat} command is needed to register
-new-style datatypes for use with \keyw{fun} and \keyw{function}
+The next example is defined using \keyw{fun} to escape the syntactic
+restrictions imposed on primitive recursive functions. The
+@{command datatype_new_compat} command is needed to register new-style datatypes
+for use with \keyw{fun} and \keyw{function}
 (Section~\ref{sssec:datatype-new-compat}):
 *}
 
@@ -1124,28 +1151,51 @@
 (@{text \<Rightarrow>}) is simply composition (@{text "op \<circ>"}):
 *}
 
-    primrec_new (*<*)(in early) (*>*)ftree_map :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
-      "ftree_map f (FTLeaf x) = FTLeaf (f x)" |
-      "ftree_map f (FTNode g) = FTNode (ftree_map f \<circ> g)"
+    primrec_new (*<*)(in early) (*>*)relabel_ft :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
+      "relabel_ft f (FTLeaf x) = FTLeaf (f x)" |
+      "relabel_ft f (FTNode g) = FTNode (relabel_ft f \<circ> g)"
+
+text {*
+\noindent
+For convenience, recursion through functions can also be expressed using
+$\lambda$-abstractions and function application rather than through composition.
+For example:
+*}
+
+    primrec_new relabel_ft :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
+      "relabel_ft f (FTLeaf x) = FTLeaf (f x)" |
+      "relabel_ft f (FTNode g) = FTNode (\<lambda>x. relabel_ft f (g x))"
+
+text {* \blankline *}
+
+    primrec_new subtree_ft :: "'a \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
+      "subtree_ft x (FTNode g) = g x"
 
 text {*
 \noindent
-(No such map function is defined by the package because the type
-variable @{typ 'a} is dead in @{typ "'a ftree"}.)
-
-Using \keyw{fun} or \keyw{function}, recursion through functions can be
-expressed using $\lambda$-expressions and function application rather
-than through composition. For example:
+For recursion through curried $n$-ary functions, $n$ applications of
+@{term "op \<circ>"} are necessary. The examples below illustrate the case where
+$n = 2$:
 *}
 
-    datatype_new_compat ftree
+    datatype_new 'a ftree2 = FTLeaf2 'a | FTNode2 "'a \<Rightarrow> 'a \<Rightarrow> 'a ftree2"
 
 text {* \blankline *}
 
-    function ftree_map :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
-      "ftree_map f (FTLeaf x) = FTLeaf (f x)" |
-      "ftree_map f (FTNode g) = FTNode (\<lambda>x. ftree_map f (g x))"
-    by auto (metis ftree.exhaust)
+    primrec_new (*<*)(in early) (*>*)relabel_ft2 :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree2 \<Rightarrow> 'a ftree2" where
+      "relabel_ft2 f (FTLeaf2 x) = FTLeaf2 (f x)" |
+      "relabel_ft2 f (FTNode2 g) = FTNode2 (op \<circ> (op \<circ> (relabel_ft2 f)) g)"
+
+text {* \blankline *}
+
+    primrec_new relabel_ft2 :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree2 \<Rightarrow> 'a ftree2" where
+      "relabel_ft2 f (FTLeaf2 x) = FTLeaf2 (f x)" |
+      "relabel_ft2 f (FTNode2 g) = FTNode2 (\<lambda>x y. relabel_ft2 f (g x y))"
+
+text {* \blankline *}
+
+    primrec_new subtree_ft2 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a ftree2 \<Rightarrow> 'a ftree2" where
+      "subtree_ft2 x y (FTNode2 g) = g x y"
 
 
 subsubsection {* Nested-as-Mutual Recursion
@@ -1177,12 +1227,12 @@
 
 text {*
 \noindent
-Appropriate induction principles are generated under the names
+Appropriate induction rules are generated as
 @{thm [source] at\<^sub>f\<^sub>f.induct},
 @{thm [source] ats\<^sub>f\<^sub>f.induct}, and
-@{thm [source] at\<^sub>f\<^sub>f_ats\<^sub>f\<^sub>f.induct}.
-
-%%% TODO: Add recursors.
+@{thm [source] at\<^sub>f\<^sub>f_ats\<^sub>f\<^sub>f.induct}. The
+induction rules and the underlying recursors are generated on a per-need basis
+and are kept in a cache to speed up subsequent definitions.
 
 Here is a second example:
 *}
@@ -1340,7 +1390,7 @@
 \begin{itemize}
 \setlength{\itemsep}{0pt}
 
-\item \emph{Theorems sometimes have different names.}
+\item \emph{Some theorems have different names.}
 For $m > 1$ mutually recursive functions,
 @{text "f\<^sub>1_\<dots>_f\<^sub>m.simps"} has been broken down into separate
 subcollections @{text "f\<^sub>i.simps"}.
@@ -1415,7 +1465,7 @@
 text {*
 \noindent
 Notice that the @{const cont} selector is associated with both @{const Skip}
-and @{const Choice}.
+and @{const Action}.
 *}
 
 
@@ -1606,10 +1656,10 @@
   \label{sec:defining-corecursive-functions} *}
 
 text {*
-Corecursive functions can be specified using @{command primcorec} and
-@{command primcorecursive}, which support primitive corecursion, or using the
-more general \keyw{partial\_function} command. Here, the focus is on
-the former two. More examples can be found in the directory
+Corecursive functions can be specified using the @{command primcorec} and
+\keyw{prim\-corec\-ursive} commands, which support primitive corecursion, or
+using the more general \keyw{partial\_function} command. Here, the focus is on
+the first two. More examples can be found in the directory
 \verb|~~/src/HOL/BNF/Examples|.
 
 Whereas recursive functions consume datatypes one constructor at a time,
@@ -1630,7 +1680,7 @@
 This style is popular in the coalgebraic literature.
 
 \item The \emph{constructor view} specifies $f$ by equations of the form
-\[@{text "\<dots> \<Longrightarrow> f x\<^sub>1 \<dots> x\<^sub>n = C \<dots>"}\]
+\[@{text "\<dots> \<Longrightarrow> f x\<^sub>1 \<dots> x\<^sub>n = C\<^sub>j \<dots>"}\]
 This style is often more concise than the previous one.
 
 \item The \emph{code view} specifies $f$ by a single equation of the form
@@ -1643,14 +1693,6 @@
 All three styles are available as input syntax. Whichever syntax is chosen,
 characteristic theorems for all three styles are generated.
 
-\begin{framed}
-\noindent
-\textbf{Warning:}\enskip The @{command primcorec} and @{command primcorecursive}
-commands are under development. Some of the functionality described here is
-vaporware. An alternative is to define corecursive functions directly using the
-generated @{text t_unfold} or @{text t_corec} combinators.
-\end{framed}
-
 %%% TODO: partial_function? E.g. for defining tail recursive function on lazy
 %%% lists (cf. terminal0 in TLList.thy)
 *}
@@ -1668,11 +1710,6 @@
 present the same examples expressed using the constructor and destructor views.
 *}
 
-(*<*)
-    locale code_view
-    begin
-(*>*)
-
 subsubsection {* Simple Corecursion
   \label{sssec:primcorec-simple-corecursion} *}
 
@@ -1683,19 +1720,19 @@
 *}
 
     primcorec literate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist" where
-      "literate f x = LCons x (literate f (f x))"
+      "literate g x = LCons x (literate g (g x))"
 
 text {* \blankline *}
 
     primcorec siterate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a stream" where
-      "siterate f x = SCons x (siterate f (f x))"
+      "siterate g x = SCons x (siterate g (g x))"
 
 text {*
 \noindent
 The constructor ensures that progress is made---i.e., the function is
 \emph{productive}. The above functions compute the infinite lazy list or stream
-@{text "[x, f x, f (f x), \<dots>]"}. Productivity guarantees that prefixes
-@{text "[x, f x, f (f x), \<dots>, (f ^^ k) x]"} of arbitrary finite length
+@{text "[x, g x, g (g x), \<dots>]"}. Productivity guarantees that prefixes
+@{text "[x, g x, g (g x), \<dots>, (g ^^ k) x]"} of arbitrary finite length
 @{text k} can be computed by unfolding the code equation a finite number of
 times.
 
@@ -1714,7 +1751,7 @@
 appear around constructors that guard corecursive calls:
 *}
 
-    primcorec_notyet lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
+    primcorec lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
       "lappend xs ys =
          (case xs of
             LNil \<Rightarrow> ys
@@ -1735,7 +1772,7 @@
 pseudorandom seed (@{text n}):
 *}
 
-    primcorec_notyet
+    primcorec
       random_process :: "'a stream \<Rightarrow> (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> 'a process"
     where
       "random_process s f n =
@@ -1780,43 +1817,71 @@
 The next pair of examples generalize the @{const literate} and @{const siterate}
 functions (Section~\ref{sssec:primcorec-nested-corecursion}) to possibly
 infinite trees in which subnodes are organized either as a lazy list (@{text
-tree\<^sub>i\<^sub>i}) or as a finite set (@{text tree\<^sub>i\<^sub>s}):
+tree\<^sub>i\<^sub>i}) or as a finite set (@{text tree\<^sub>i\<^sub>s}). They rely on the map functions of
+the nesting type constructors to lift the corecursive calls:
 *}
 
     primcorec iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
-      "iterate\<^sub>i\<^sub>i f x = Node\<^sub>i\<^sub>i x (lmap (iterate\<^sub>i\<^sub>i f) (f x))"
+      "iterate\<^sub>i\<^sub>i g x = Node\<^sub>i\<^sub>i x (lmap (iterate\<^sub>i\<^sub>i g) (g x))"
 
 text {* \blankline *}
 
     primcorec iterate\<^sub>i\<^sub>s :: "('a \<Rightarrow> 'a fset) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>s" where
-      "iterate\<^sub>i\<^sub>s f x = Node\<^sub>i\<^sub>s x (fimage (iterate\<^sub>i\<^sub>s f) (f x))"
+      "iterate\<^sub>i\<^sub>s g x = Node\<^sub>i\<^sub>s x (fimage (iterate\<^sub>i\<^sub>s g) (g x))"
 
 text {*
 \noindent
-Deterministic finite automata (DFAs) are traditionally defined as 5-tuples
-@{text "(Q, \<Sigma>, \<delta>, q\<^sub>0, F)"}, where @{text Q} is a finite set of states,
+Both examples follow the usual format for constructor arguments associated
+with nested recursive occurrences of the datatype. Consider
+@{const iterate\<^sub>i\<^sub>i}. The term @{term "g x"} constructs an @{typ "'a llist"}
+value, which is turned into an @{typ "'a tree\<^sub>i\<^sub>i llist"} value using
+@{const lmap}.
+
+This format may sometimes feel artificial. The following function constructs
+a tree with a single, infinite branch from a stream:
+*}
+
+    primcorec tree\<^sub>i\<^sub>i_of_stream :: "'a stream \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
+      "tree\<^sub>i\<^sub>i_of_stream s =
+         Node\<^sub>i\<^sub>i (shd s) (lmap tree\<^sub>i\<^sub>i_of_stream (LCons (stl s) LNil))"
+
+text {*
+\noindent
+Fortunately, it is easy to prove the following lemma, where the corecursive call
+is moved inside the lazy list constructor, thereby eliminating the need for
+@{const lmap}:
+*}
+
+    lemma tree\<^sub>i\<^sub>i_of_stream_alt:
+      "tree\<^sub>i\<^sub>i_of_stream s = Node\<^sub>i\<^sub>i (shd s) (LCons (tree\<^sub>i\<^sub>i_of_stream (stl s)) LNil)"
+    by (subst tree\<^sub>i\<^sub>i_of_stream.code) simp
+
+text {*
+The next example illustrates corecursion through functions, which is a bit
+special. Deterministic finite automata (DFAs) are traditionally defined as
+5-tuples @{text "(Q, \<Sigma>, \<delta>, q\<^sub>0, F)"}, where @{text Q} is a finite set of states,
 @{text \<Sigma>} is a finite alphabet, @{text \<delta>} is a transition function, @{text q\<^sub>0}
 is an initial state, and @{text F} is a set of final states. The following
 function translates a DFA into a @{type state_machine}:
 *}
 
-    primcorec (*<*)(in early) (*>*)
-      sm_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> 'a state_machine"
+    primcorec
+      (*<*)(in early) (*>*)sm_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> 'a state_machine"
     where
-      "sm_of_dfa \<delta> F q = State_Machine (q \<in> F) (sm_of_dfa \<delta> F o \<delta> q)"
+      "sm_of_dfa \<delta> F q = State_Machine (q \<in> F) (sm_of_dfa \<delta> F \<circ> \<delta> q)"
 
 text {*
 \noindent
 The map function for the function type (@{text \<Rightarrow>}) is composition
-(@{text "op \<circ>"}). For convenience, corecursion through functions can be
-expressed using $\lambda$-expressions and function application rather
+(@{text "op \<circ>"}). For convenience, corecursion through functions can
+also be expressed using $\lambda$-abstractions and function application rather
 than through composition. For example:
 *}
 
     primcorec
       sm_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> 'a state_machine"
     where
-      "sm_of_dfa \<delta> F q = State_Machine (q \<in> F) (sm_of_dfa \<delta> F o \<delta> q)"
+      "sm_of_dfa \<delta> F q = State_Machine (q \<in> F) (\<lambda>a. sm_of_dfa \<delta> F (\<delta> q a))"
 
 text {* \blankline *}
 
@@ -1833,9 +1898,32 @@
     primcorec
       or_sm :: "'a state_machine \<Rightarrow> 'a state_machine \<Rightarrow> 'a state_machine"
     where
-      "or_sm M N =
-         State_Machine (accept M \<or> accept N)
-           (\<lambda>a. or_sm (trans M a) (trans N a))"
+      "or_sm M N = State_Machine (accept M \<or> accept N)
+         (\<lambda>a. or_sm (trans M a) (trans N a))"
+
+text {*
+\noindent
+For recursion through curried $n$-ary functions, $n$ applications of
+@{term "op \<circ>"} are necessary. The examples below illustrate the case where
+$n = 2$:
+*}
+
+    codatatype ('a, 'b) state_machine2 =
+      State_Machine2 (accept2: bool) (trans2: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) state_machine2")
+
+text {* \blankline *}
+
+    primcorec
+      (*<*)(in early) (*>*)sm2_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> ('a, 'b) state_machine2"
+    where
+      "sm2_of_dfa \<delta> F q = State_Machine2 (q \<in> F) (op \<circ> (op \<circ> (sm2_of_dfa \<delta> F)) (\<delta> q))"
+
+text {* \blankline *}
+
+    primcorec
+      sm2_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> ('a, 'b) state_machine2"
+    where
+      "sm2_of_dfa \<delta> F q = State_Machine2 (q \<in> F) (\<lambda>a b. sm2_of_dfa \<delta> F (\<delta> q a b))"
 
 
 subsubsection {* Nested-as-Mutual Corecursion
@@ -1848,15 +1936,31 @@
 pretend that nested codatatypes are mutually corecursive. For example:
 *}
 
-    primcorec_notyet
+(*<*)
+    context late
+    begin
+(*>*)
+    primcorec
       iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" and
       iterates\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a llist \<Rightarrow> 'a tree\<^sub>i\<^sub>i llist"
     where
-      "iterate\<^sub>i\<^sub>i f x = Node\<^sub>i\<^sub>i x (iterates\<^sub>i\<^sub>i f (f x))" |
-      "iterates\<^sub>i\<^sub>i f xs =
+      "iterate\<^sub>i\<^sub>i g x = Node\<^sub>i\<^sub>i x (iterates\<^sub>i\<^sub>i g (g x))" |
+      "iterates\<^sub>i\<^sub>i g xs =
          (case xs of
             LNil \<Rightarrow> LNil
-          | LCons x xs' \<Rightarrow> LCons (iterate\<^sub>i\<^sub>i f x) (iterates\<^sub>i\<^sub>i f xs'))"
+          | LCons x xs' \<Rightarrow> LCons (iterate\<^sub>i\<^sub>i g x) (iterates\<^sub>i\<^sub>i g xs'))"
+
+text {*
+\noindent
+Coinduction rules are generated as
+@{thm [source] iterate\<^sub>i\<^sub>i.coinduct},
+@{thm [source] iterates\<^sub>i\<^sub>i.coinduct}, and
+@{thm [source] iterate\<^sub>i\<^sub>i_iterates\<^sub>i\<^sub>i.coinduct}
+and analogously for @{text strong_coinduct}. These rules and the
+underlying corecursors are generated on a per-need basis and are kept in a cache
+to speed up subsequent definitions.
+*}
+
 (*<*)
     end
 (*>*)
@@ -1866,7 +1970,7 @@
   \label{ssec:primrec-constructor-view} *}
 
 (*<*)
-    locale ctr_view = code_view
+    locale ctr_view
     begin
 (*>*)
 
@@ -1937,7 +2041,7 @@
   \label{ssec:primrec-destructor-view} *}
 
 (*<*)
-    locale dest_view
+    locale dtr_view
     begin
 (*>*)
 
@@ -1951,13 +2055,13 @@
     primcorec literate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist" where
       "\<not> lnull (literate _ x)" |
       "lhd (literate _ x) = x" |
-      "ltl (literate f x) = literate f (f x)"
+      "ltl (literate g x) = literate g (g x)"
 
 text {* \blankline *}
 
     primcorec siterate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a stream" where
       "shd (siterate _ x) = x" |
-      "stl (siterate f x) = siterate f (f x)"
+      "stl (siterate g x) = siterate g (g x)"
 
 text {* \blankline *}
 
@@ -1993,6 +2097,9 @@
 (*<*)
     end
 
+    locale dtr_view2
+    begin
+
     primcorec lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
       "lnull xs \<Longrightarrow> lnull ys \<Longrightarrow> lnull (lappend xs ys)" |
 (*>*)
@@ -2000,8 +2107,6 @@
 (*<*) |
       "lhd (lappend xs ys) = lhd (if lnull xs then ys else xs)" |
       "ltl (lappend xs ys) = (if xs = LNil then ltl ys else lappend (ltl xs) ys)"
-
-    context dest_view begin
 (*>*)
 
 text {*
@@ -2044,8 +2149,8 @@
 text {* \blankline *}
 
     primcorec iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
-      "lbl\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i f x) = x" |
-      "sub\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i f x) = lmap (iterate\<^sub>i\<^sub>i f) (f x)"
+      "lbl\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i g x) = x" |
+      "sub\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i g x) = lmap (iterate\<^sub>i\<^sub>i g) (g x)"
 (*<*)
     end
 (*>*)
@@ -2149,13 +2254,30 @@
 
 @{rail "
   @@{command bnf} target? (name ':')? term \\
-    term_list term term_list term?
+    term_list term term_list? term?
   ;
   X_list: '[' (X + ',') ']'
 "}
 *}
 
 
+(* NOTYET
+subsubsection {* \keyw{bnf\_decl}
+  \label{sssec:bnf-decl} *}
+
+text {*
+%%% TODO: use command_def once the command is available
+\begin{matharray}{rcl}
+  @{text "bnf_decl"} & : & @{text "local_theory \<rightarrow> local_theory"}
+\end{matharray}
+
+@{rail "
+  @@{command bnf} target? dt_name
+"}
+*}
+*)
+
+
 subsubsection {* \keyw{print\_bnfs}
   \label{sssec:print-bnfs} *}
 
@@ -2307,8 +2429,9 @@
 suggested major simplifications to the internal constructions, much of which has
 yet to be implemented. Florian Haftmann and Christian Urban provided general
 advice on Isabelle and package writing. Stefan Milius and Lutz Schr\"oder
-found an elegant proof to eliminate one of the BNF assumptions. Christian
-Sternagel suggested many textual improvements to this tutorial.
+found an elegant proof to eliminate one of the BNF assumptions. Andreas
+Lochbihler and Christian Sternagel suggested many textual improvements to this
+tutorial.
 *}
 
 end
--- a/src/Doc/Datatypes/document/root.tex	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/Doc/Datatypes/document/root.tex	Mon Nov 11 17:44:21 2013 +0100
@@ -58,10 +58,10 @@
 
 \begin{abstract}
 \noindent
-This tutorial describes how to use the new package for defining datatypes and
-codatatypes in Isabelle/HOL. The package provides five main commands:
+This tutorial describes the new package for defining datatypes and codatatypes
+in Isabelle/HOL. The package provides four main commands:
 \keyw{datatype\_new}, \keyw{codatatype}, \keyw{primrec\_new},
-\keyw{primcorecursive}, and \keyw{primcorec}. The commands suffixed by
+and \keyw{primcorec}. The commands suffixed by
 \keyw{\_new} are intended to subsume, and eventually replace, the corresponding
 commands from the old datatype package.
 \end{abstract}
--- a/src/Doc/Functions/Functions.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/Doc/Functions/Functions.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1003,13 +1003,13 @@
   recursive calls. In general, there is one introduction rule for each
   recursive call.
 
-  The predicate @{term "accp findzero_rel"} is the accessible part of
+  The predicate @{term "Wellfounded.accp findzero_rel"} is the accessible part of
   that relation. An argument belongs to the accessible part, if it can
   be reached in a finite number of steps (cf.~its definition in @{text
   "Wellfounded.thy"}).
 
   Since the domain predicate is just an abbreviation, you can use
-  lemmas for @{const accp} and @{const findzero_rel} directly. Some
+  lemmas for @{const Wellfounded.accp} and @{const findzero_rel} directly. Some
   lemmas which are occasionally useful are @{thm [source] accpI}, @{thm [source]
   accp_downward}, and of course the introduction and elimination rules
   for the recursion relation @{thm [source] "findzero_rel.intros"} and @{thm
--- a/src/Doc/Nitpick/document/root.tex	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/Doc/Nitpick/document/root.tex	Mon Nov 11 17:44:21 2013 +0100
@@ -1965,6 +1965,8 @@
 \texttt{.kki}, \texttt{.cnf}, \texttt{.out}, and
 \texttt{.err}; you may safely remove them after Nitpick has run.
 
+\textbf{Warning:} This option is not thread-safe. Use at your own risks.
+
 \nopagebreak
 {\small See also \textit{debug} (\S\ref{output-format}).}
 \end{enum}
@@ -2794,11 +2796,12 @@
 \subsection{Registering Coinductive Datatypes}
 \label{registering-coinductive-datatypes}
 
+Coinductive datatypes defined using the \textbf{codatatype} command that do not
+involve nested recursion through non-codatatypes are supported by Nitpick.
 If you have defined a custom coinductive datatype, you can tell Nitpick about
-it, so that it can use an efficient Kodkod axiomatization similar to the one it
-uses for lazy lists. The interface for registering and unregistering coinductive
-datatypes consists of the following pair of functions defined in the
-\textit{Nitpick\_HOL} structure:
+it, so that it can use an efficient Kodkod axiomatization. The interface for
+registering and unregistering coinductive datatypes consists of the following
+pair of functions defined in the \textit{Nitpick\_HOL} structure:
 
 \prew
 $\textbf{val}\,~\textit{register\_codatatype\/} : {}$ \\
@@ -2886,6 +2889,12 @@
 \item[\labelitemi] Nitpick produces spurious counterexamples when invoked after a
 \textbf{guess} command in a structured proof.
 
+\item[\labelitemi] Datatypes defined using \textbf{datatype\_new} are not
+supported.
+
+\item[\labelitemi] Codatatypes defined using \textbf{codatatype} that
+involve nested recursion through non-codatatypes are not supported.
+
 \item[\labelitemi] The \textit{nitpick\_xxx} attributes and the
 \textit{Nitpick\_xxx.register\_yyy} functions can cause havoc if used
 improperly.
--- a/src/Doc/ProgProve/Bool_nat_list.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/Doc/ProgProve/Bool_nat_list.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -422,10 +422,16 @@
 \subsection{Exercises}
 
 \begin{exercise}
+Use the \isacom{value} command to evaluate the following expressions:
+@{term[source] "1 + (2::nat)"}, @{term[source] "1 + (2::int)"},
+@{term[source] "1 - (2::nat)"} and @{term[source] "1 - (2::int)"}.
+\end{exercise}
+
+\begin{exercise}
 Start from the definition of @{const add} given above.
-Prove it is associative (@{prop"add (add m n) p = add m (add n p)"})
-and commutative (@{prop"add m n = add n m"}). Define a recursive function
-@{text double} @{text"::"} @{typ"nat \<Rightarrow> nat"} and prove that @{prop"double m = add m m"}.
+Prove that @{const add} it is associative and commutative.
+Define a recursive function @{text double} @{text"::"} @{typ"nat \<Rightarrow> nat"}
+and prove @{prop"double m = add m m"}.
 \end{exercise}
 
 \begin{exercise}
@@ -436,11 +442,15 @@
 
 \begin{exercise}
 Define a recursive function @{text "snoc ::"} @{typ"'a list \<Rightarrow> 'a \<Rightarrow> 'a list"}
-that appends an element to the end of a list. Do not use the predefined append
-operator @{text"@"}. With the help of @{text snoc} define a recursive function
-@{text "reverse ::"} @{typ"'a list \<Rightarrow> 'a list"} that reverses a list. Do not
-use the predefined function @{const rev}.
-Prove @{prop"reverse(reverse xs) = xs"}.
+that appends an element to the end of a list. With the help of @{text snoc}
+define a recursive function @{text "reverse ::"} @{typ"'a list \<Rightarrow> 'a list"}
+that reverses a list. Prove @{prop"reverse(reverse xs) = xs"}.
+\end{exercise}
+
+\begin{exercise}
+Define a recursive function @{text "sum ::"} @{typ"nat \<Rightarrow> nat"} such that
+\mbox{@{text"sum n"}} @{text"="} @{text"0 + ... + n"} and prove
+@{prop" sum(n::nat) = n * (n+1) div 2"}.
 \end{exercise}
 *}
 (*<*)
--- a/src/Doc/ProgProve/Isar.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/Doc/ProgProve/Isar.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -595,10 +595,10 @@
 \exercise
 Give a readable, structured proof of the following lemma:
 *}
-lemma assumes T: "\<forall> x y. T x y \<or> T y x"
-  and A: "\<forall> x y. A x y \<and> A y x \<longrightarrow> x = y"
-  and TA: "\<forall> x y. T x y \<longrightarrow> A x y" and "A x y"
-shows "T x y"
+lemma assumes T: "\<forall>x y. T x y \<or> T y x"
+  and A: "\<forall>x y. A x y \<and> A y x \<longrightarrow> x = y"
+  and TA: "\<forall>x y. T x y \<longrightarrow> A x y" and "A x y"
+  shows "T x y"
 (*<*)oops(*>*)
 text{*
 \endexercise
@@ -612,10 +612,11 @@
 text{*
 Hint: There are predefined functions @{const_typ take} and @{const_typ drop}
 such that @{text"take k [x\<^sub>1,\<dots>] = [x\<^sub>1,\<dots>,x\<^sub>k]"} and
-@{text"drop k [x\<^sub>1,\<dots>] = [x\<^bsub>k+1\<^esub>,\<dots>]"}. Let @{text simp} and especially
-sledgehammer find and apply the relevant @{const take} and @{const drop} lemmas for you.
+@{text"drop k [x\<^sub>1,\<dots>] = [x\<^bsub>k+1\<^esub>,\<dots>]"}. Let sledgehammer find and apply
+the relevant @{const take} and @{const drop} lemmas for you.
 \endexercise
 
+
 \section{Case Analysis and Induction}
 
 \subsection{Datatype Case Analysis}
@@ -1075,45 +1076,38 @@
 @{text induct} method.
 \end{warn}
 
+
 \subsection{Exercises}
 
+
+\exercise
+Give a structured proof by rule inversion:
+*}
+
+lemma assumes a: "ev(Suc(Suc n))" shows "ev n"
+(*<*)oops(*>*)
+
+text{*
+\endexercise
+
+\begin{exercise}
+Give a structured proof of @{prop "\<not> ev(Suc(Suc(Suc 0)))"}
+by rule inversions. If there are no cases to be proved you can close
+a proof immediateley with \isacom{qed}.
+\end{exercise}
+
+\begin{exercise}
+Recall predicate @{text star} from \autoref{sec:star} and @{text iter}
+from Exercise~\ref{exe:iter}. Prove @{prop "iter r n x y \<Longrightarrow> star r x y"}
+in a structured style, do not just sledgehammer each case of the
+required induction.
+\end{exercise}
+
 \begin{exercise}
 Define a recursive function @{text "elems ::"} @{typ"'a list \<Rightarrow> 'a set"}
 and prove @{prop "x : elems xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> elems ys"}.
 \end{exercise}
-
-\begin{exercise}
-A context-free grammar can be seen as an inductive definition where each
-nonterminal $A$ is an inductively defined predicate on lists of terminal
-symbols: $A(w)$ mans
-that $w$ is in the language generated by $A$. For example, the production $S
-\to a S b$ can be viewed as the implication @{prop"S w \<Longrightarrow> S (a # w @ [b])"}
-where @{text a} and @{text b} are constructors of some datatype of terminal
-symbols: \isacom{datatype} @{text"tsymbs = a | b | \<dots>"}
-
-Define the two grammars
-\[
-\begin{array}{r@ {\quad}c@ {\quad}l}
-S &\to& \varepsilon \quad\mid\quad a~S~b \quad\mid\quad S~S \\
-T &\to& \varepsilon \quad\mid\quad T~a~T~b
-\end{array}
-\]
-($\varepsilon$ is the empty word)
-as two inductive predicates and prove @{prop"S w \<longleftrightarrow> T w"}.
-\end{exercise}
-
 *}
-(*
-lemma "\<not> ev(Suc(Suc(Suc 0)))"
-proof
-  assume "ev(Suc(Suc(Suc 0)))"
-  then show False
-  proof cases
-    case evSS
-    from `ev(Suc 0)` show False by cases
-  qed
-qed
-*)
 
 (*<*)
 end
--- a/src/Doc/ProgProve/Logic.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/Doc/ProgProve/Logic.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -141,6 +141,28 @@
 See \cite{Nipkow-Main} for the wealth of further predefined functions in theory
 @{theory Main}.
 
+
+\subsection{Exercises}
+
+\exercise
+Start from the data type of binary trees defined earlier:
+*}
+
+datatype 'a tree = Tip | Node "'a tree" 'a "'a tree"
+
+text{*
+Define a function @{text "set ::"} @{typ "'a tree \<Rightarrow> 'a set"}
+that returns the elements in a tree and a function
+@{text "ord ::"} @{typ "int tree \<Rightarrow> bool"}
+the tests if an @{typ "int tree"} is ordered.
+
+Define a function @{text ins} that inserts an element into an ordered @{typ "int tree"}
+while maintaining the order of the tree. If the element is already in the tree, the
+same tree should be returned. Prove correctness of @{text ins}:
+@{prop "set(ins x t) = {x} \<union> set t"} and @{prop "ord t \<Longrightarrow> ord(ins i t)"}.
+\endexercise
+
+
 \section{Proof Automation}
 
 So far we have only seen @{text simp} and @{text auto}: Both perform
@@ -459,12 +481,12 @@
 text{* In this particular example we could have backchained with
 @{thm[source] Suc_leD}, too, but because the premise is more complicated than the conclusion this can easily lead to nontermination.
 
-\subsection{Finding Theorems}
-
-Command \isacom{find{\isacharunderscorekeyword}theorems} searches for specific theorems in the current
-theory. Search criteria include pattern matching on terms and on names.
-For details see the Isabelle/Isar Reference Manual~\cite{IsarRef}.
-\bigskip
+%\subsection{Finding Theorems}
+%
+%Command \isacom{find{\isacharunderscorekeyword}theorems} searches for specific theorems in the current
+%theory. Search criteria include pattern matching on terms and on names.
+%For details see the Isabelle/Isar Reference Manual~\cite{IsarRef}.
+%\bigskip
 
 \begin{warn}
 To ease readability we will drop the question marks
@@ -708,8 +730,8 @@
 apply(rename_tac u x y)
 defer
 (*>*)
-txt{* The induction is over @{prop"star r x y"} and we try to prove
-\mbox{@{prop"star r y z \<Longrightarrow> star r x z"}},
+txt{* The induction is over @{prop"star r x y"} (the first matching assumption)
+and we try to prove \mbox{@{prop"star r y z \<Longrightarrow> star r x z"}},
 which we abbreviate by @{prop"P x y"}. These are our two subgoals:
 @{subgoals[display,indent=0]}
 The first one is @{prop"P x x"}, the result of case @{thm[source]refl},
@@ -764,6 +786,95 @@
 conditions}. In rule inductions, these side-conditions appear as additional
 assumptions. The \isacom{for} clause seen in the definition of the reflexive
 transitive closure merely simplifies the form of the induction rule.
+
+
+\subsection{Exercises}
+
+\begin{exercise}
+Formalise the following definition of palindromes
+\begin{itemize}
+\item The empty list and a singleton list are palindromes.
+\item If @{text xs} is a palindrome, so is @{term "a # xs @ [a]"}.
+\end{itemize}
+as an inductive predicate @{text "palindrome ::"} @{typ "'a list \<Rightarrow> bool"}
+and prove that @{prop "rev xs = xs"} if @{text xs} is a palindrome.
+\end{exercise}
+
+\exercise
+We could also have defined @{const star} as follows:
+*}
+
+inductive star' :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" for r where
+refl': "star' r x x" |
+step': "star' r x y \<Longrightarrow> r y z \<Longrightarrow> star' r x z"
+
+text{*
+The single @{text r} step is performer after rather than before the @{text star'}
+steps. Prove @{prop "star' r x y \<Longrightarrow> star r x y"} and
+@{prop "star r x y \<Longrightarrow> star r' x y"}. You may need lemmas.
+Note that rule induction fails
+if the assumption about the inductive predicate is not the first assumption.
+\endexercise
+
+\begin{exercise}\label{exe:iter}
+Analogous to @{const star}, give an inductive definition of the @{text n}-fold iteration
+of a relation @{text r}: @{term "iter r n x y"} should hold if there are @{text x\<^sub>0}, \dots, @{text x\<^sub>n}
+such that @{prop"x = x\<^sub>0"}, @{prop"x\<^sub>n = y"} and @{text"r x\<^bsub>i\<^esub> x\<^bsub>i+1\<^esub>"} for
+all @{prop"i < n"}. Correct and prove the following claim:
+@{prop"star r x y \<Longrightarrow> iter r n x y"}.
+\end{exercise}
+
+\begin{exercise}
+A context-free grammar can be seen as an inductive definition where each
+nonterminal $A$ is an inductively defined predicate on lists of terminal
+symbols: $A(w)$ mans that $w$ is in the language generated by $A$.
+For example, the production $S \to a S b$ can be viewed as the implication
+@{prop"S w \<Longrightarrow> S (a # w @ [b])"} where @{text a} and @{text b} are terminal symbols,
+i.e., elements of some alphabet. The alphabet can be defined like this:
+\isacom{datatype} @{text"alpha = a | b | \<dots>"}
+
+Define the two grammars (where $\varepsilon$ is the empty word)
+\[
+\begin{array}{r@ {\quad}c@ {\quad}l}
+S &\to& \varepsilon \quad\mid\quad aSb \quad\mid\quad SS \\
+T &\to& \varepsilon \quad\mid\quad TaTb
+\end{array}
+\]
+as two inductive predicates.
+If you think of @{text a} and @{text b} as ``@{text "("}'' and  ``@{text ")"}'',
+the grammars defines strings of balanced parentheses.
+Prove @{prop"T w \<Longrightarrow> S w"} and @{prop "S w \<Longrightarrow> T w"} separately and conclude
+@{prop "S w = T w"}.
+\end{exercise}
+
+\ifsem
+\begin{exercise}
+In \autoref{sec:AExp} we defined a recursive evaluation function
+@{text "aval :: aexp \<Rightarrow> state \<Rightarrow> val"}.
+Define an inductive evaluation predicate
+@{text "aval_rel :: aexp \<Rightarrow> state \<Rightarrow> val \<Rightarrow> bool"}
+and prove that it agrees with the recursive function:
+@{prop "aval_rel a s v \<Longrightarrow> aval a s = v"}, 
+@{prop "aval a s = v \<Longrightarrow> aval_rel a s v"} and thus
+\noquotes{@{prop [source] "aval_rel a s v \<longleftrightarrow> aval a s = v"}}.
+\end{exercise}
+
+\begin{exercise}
+Consider the stack machine from Chapter~3
+and recall the concept of \concept{stack underflow}
+from Exercise~\ref{exe:stack-underflow}.
+Define an inductive predicate
+@{text "ok :: nat \<Rightarrow> instr list \<Rightarrow> nat \<Rightarrow> bool"}
+such that @{text "ok n is n'"} means that with any initial stack of length
+@{text n} the instructions @{text "is"} can be executed
+without stack underflow and that the final stack has length @{text n'}.
+Prove that @{text ok} correctly computes the final stack size
+@{prop[display] "\<lbrakk>ok n is n'; length stk = n\<rbrakk> \<Longrightarrow> length (exec is s stk) = n'"}
+and that instruction sequences generated by @{text comp}
+cannot cause stack underflow: \ @{text "ok n (comp a) ?"} \ for
+some suitable value of @{text "?"}.
+\end{exercise}
+\fi
 *}
 (*<*)
 end
--- a/src/Doc/ProgProve/Types_and_funs.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/Doc/ProgProve/Types_and_funs.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -156,7 +156,7 @@
 
 fun div2 :: "nat \<Rightarrow> nat" where
 "div2 0 = 0" |
-"div2 (Suc 0) = Suc 0" |
+"div2 (Suc 0) = 0" |
 "div2 (Suc(Suc n)) = Suc(div2 n)"
 
 text{* does not just define @{const div2} but also proves a
@@ -200,6 +200,34 @@
 But note that the induction rule does not mention @{text f} at all,
 except in its name, and is applicable independently of @{text f}.
 
+
+\subsection{Exercises}
+
+\begin{exercise}
+Starting from the type @{text "'a tree"} defined in the text, define
+a function @{text "contents ::"} @{typ "'a tree \<Rightarrow> 'a list"}
+that collects all values in a tree in a list, in any order,
+without removing duplicates.
+Then define a function @{text "treesum ::"} @{typ "nat tree \<Rightarrow> nat"}
+that sums up all values in a tree of natural numbers
+and prove @{prop "treesum t = listsum(contents t)"}.
+\end{exercise}
+
+\begin{exercise}
+Define a new type @{text "'a tree2"} of binary trees where values are also
+stored in the leaves of the tree.  Also reformulate the
+@{const mirror} function accordingly. Define two functions
+@{text "pre_order"} and @{text "post_order"} of type @{text "'a tree2 \<Rightarrow> 'a list"}
+that traverse a tree and collect all stored values in the respective order in
+a list. Prove @{prop "pre_order (mirror t) = rev (post_order t)"}.
+\end{exercise}
+
+\begin{exercise}
+Prove that @{const div2} defined above divides every number by @{text 2},
+not just those of the form @{text"n+n"}: @{prop "div2 n = n div 2"}.
+\end{exercise}
+
+
 \section{Induction Heuristics}
 
 We have already noted that theorems about recursive functions are proved by
@@ -307,6 +335,18 @@
 matters in some cases. The variables that need to be quantified are typically
 those that change in recursive calls.
 
+
+\subsection{Exercises}
+
+\begin{exercise}
+Write a tail-recursive variant of the @{text add} function on @{typ nat}:
+@{term "itadd :: nat \<Rightarrow> nat \<Rightarrow> nat"}.
+Tail-recursive means that in the recursive case, @{text itadd} needs to call
+itself directly: \mbox{@{term"itadd (Suc m) n"}} @{text"= itadd \<dots>"}.
+Prove @{prop "itadd m n = add m n"}.
+\end{exercise}
+
+
 \section{Simplification}
 
 So far we have talked a lot about simplifying terms without explaining the concept. \concept{Simplification} means
@@ -485,6 +525,31 @@
 
 \subsection{Exercises}
 
+\exercise\label{exe:tree0}
+Define a datatype @{text tree0} of binary tree skeletons which do not store
+any information, neither in the inner nodes nor in the leaves.
+Define a function @{text "nodes :: tree0 \<Rightarrow> nat"} that counts the total number
+all nodes (inner nodes and leaves) in such a tree.
+Consider the following recursive function:
+*}
+(*<*)
+datatype tree0 = Tip | Node tree0 tree0
+(*>*)
+fun explode :: "nat \<Rightarrow> tree0 \<Rightarrow> tree0" where
+"explode 0 t = t" |
+"explode (Suc n) t = explode n (Node t t)"
+
+text {*
+Find an equation expressing the size of a tree after exploding it
+(\noquotes{@{term [source] "nodes (explode n t)"}}) as a function
+of @{term "nodes t"} and @{text n}. Prove your equation.
+You may use the usual arithmetic operators including the exponentiation
+operator ``@{text"^"}''. For example, \noquotes{@{prop [source] "2 ^ 2 = 4"}}.
+
+Hint: simplifying with the list of theorems @{thm[source] algebra_simps}
+takes care of common algebraic properties of the arithmetic operators.
+\endexercise
+
 \exercise
 Define arithmetic expressions in one variable over integers (type @{typ int})
 as a data type:
@@ -506,8 +571,7 @@
 that transforms an expression into a polynomial. This may require auxiliary
 functions. Prove that @{text coeffs} preserves the value of the expression:
 \mbox{@{prop"evalp (coeffs e) x = eval e x"}.}
-Hint: simplifying with @{thm[source] algebra_simps} takes care of
-common algebraic properties of @{text "+"} and @{text "*"}.
+Hint: consider the hint in \autoref{exe:tree0}.
 \endexercise
 *}
 (*<*)
--- a/src/Doc/Sledgehammer/document/root.tex	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/Doc/Sledgehammer/document/root.tex	Mon Nov 11 17:44:21 2013 +0100
@@ -121,8 +121,8 @@
 
 For Isabelle/jEdit users, Sledgehammer provides an automatic mode that can be
 enabled via the ``Auto Sledgehammer'' option under ``Plugins > Plugin Options >
-Isabelle > General.'' In this mode, Sledgehammer is run on every newly entered
-theorem.
+Isabelle > General.'' In this mode, a reduced version of Sledgehammer is run on
+every newly entered theorem for a few seconds.
 
 \newbox\boxA
 \setbox\boxA=\hbox{\texttt{NOSPAM}}
@@ -719,12 +719,16 @@
 If you use Isabelle/jEdit, Sledgehammer also provides an automatic mode that can
 be enabled via the ``Auto Sledgehammer'' option under ``Plugins > Plugin Options
 > Isabelle > General.'' For automatic runs, only the first prover set using
-\textit{provers} (\S\ref{mode-of-operation}) is considered, fewer facts are
-passed to the prover, \textit{slice} (\S\ref{mode-of-operation}) is disabled,
-\textit{strict} (\S\ref{problem-encoding}) is enabled, \textit{verbose}
-(\S\ref{output-format}) and \textit{debug} (\S\ref{output-format}) are disabled,
-and \textit{timeout} (\S\ref{timeouts}) is superseded by the ``Auto Time Limit''
-option in jEdit. Sledgehammer's output is also more concise.
+\textit{provers} (\S\ref{mode-of-operation}) is considered (typically E),
+\textit{slice} (\S\ref{mode-of-operation}) is disabled,
+\textit{minimize} (\S\ref{mode-of-operation}) is disabled, fewer facts are
+passed to the prover, \textit{fact\_filter} (\S\ref{relevance-filter}) is set to
+\textit{mepo}, \textit{strict} (\S\ref{problem-encoding}) is enabled,
+\textit{verbose} (\S\ref{output-format}) and \textit{debug}
+(\S\ref{output-format}) are disabled, \textit{preplay\_timeout}
+(\S\ref{timeouts}) is set to 0, and \textit{timeout} (\S\ref{timeouts}) is
+superseded by the ``Auto Time Limit'' option in jEdit. Sledgehammer's output is
+also more concise.
 
 \subsection{Metis}
 
@@ -999,8 +1003,7 @@
 number of facts. For SMT solvers, several slices are tried with the same options
 each time but fewer and fewer facts. According to benchmarks with a timeout of
 30 seconds, slicing is a valuable optimization, and you should probably leave it
-enabled unless you are conducting experiments. This option is implicitly
-disabled for (short) automatic runs.
+enabled unless you are conducting experiments.
 
 \nopagebreak
 {\small See also \textit{verbose} (\S\ref{output-format}).}
@@ -1035,6 +1038,8 @@
 simultaneously. The files are identified by the prefixes \texttt{prob\_} and
 \texttt{mash\_}; you may safely remove them after Sledgehammer has run.
 
+\textbf{Warning:} This option is not thread-safe. Use at your own risks.
+
 \nopagebreak
 {\small See also \textit{debug} (\S\ref{output-format}).}
 \end{enum}
@@ -1282,14 +1287,12 @@
 
 \opfalse{verbose}{quiet}
 Specifies whether the \textbf{sledgehammer} command should explain what it does.
-This option is implicitly disabled for automatic runs.
 
 \opfalse{debug}{no\_debug}
 Specifies whether Sledgehammer should display additional debugging information
 beyond what \textit{verbose} already displays. Enabling \textit{debug} also
 enables \textit{verbose} and \textit{blocking} (\S\ref{mode-of-operation})
-behind the scenes. The \textit{debug} option is implicitly disabled for
-automatic runs.
+behind the scenes.
 
 \nopagebreak
 {\small See also \textit{spy} (\S\ref{mode-of-operation}) and
@@ -1349,8 +1352,6 @@
 \opdefault{timeout}{float\_or\_none}{\upshape 30}
 Specifies the maximum number of seconds that the automatic provers should spend
 searching for a proof. This excludes problem preparation and is a soft limit.
-For automatic runs, the ``Auto Time Limit'' option under ``Plugins > Plugin
-Options > Isabelle > General'' is used instead.
 
 \opdefault{preplay\_timeout}{float\_or\_none}{\upshape 3}
 Specifies the maximum number of seconds that \textit{metis} or \textit{smt}
--- a/src/Doc/Tutorial/document/sets.tex	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/Doc/Tutorial/document/sets.tex	Mon Nov 11 17:44:21 2013 +0100
@@ -660,8 +660,8 @@
 \textbf{Composition} of relations (the infix \sdx{O}) is also
 available: 
 \begin{isabelle}
-r\ O\ s\ \isasymequiv\ \isacharbraceleft(x,z).\ \isasymexists y.\ (x,y)\ \isasymin\ s\ \isasymand\ (y,z)\ \isasymin\ r\isacharbraceright
-\rulenamedx{rel_comp_def}
+r\ O\ s\ = \isacharbraceleft(x,z).\ \isasymexists y.\ (x,y)\ \isasymin\ s\ \isasymand\ (y,z)\ \isasymin\ r\isacharbraceright
+\rulenamedx{relcomp_unfold}
 \end{isabelle}
 %
 This is one of the many lemmas proved about these concepts: 
@@ -677,7 +677,7 @@
 \isasymlbrakk r\isacharprime\ \isasymsubseteq\ r;\ s\isacharprime\
 \isasymsubseteq\ s\isasymrbrakk\ \isasymLongrightarrow\ r\isacharprime\ O\
 s\isacharprime\ \isasymsubseteq\ r\ O\ s%
-\rulename{rel_comp_mono}
+\rulename{relcomp_mono}
 \end{isabelle}
 
 \indexbold{converse!of a relation}%
@@ -695,7 +695,7 @@
 Here is a typical law proved about converse and composition: 
 \begin{isabelle}
 (r\ O\ s)\isasyminverse\ =\ s\isasyminverse\ O\ r\isasyminverse
-\rulename{converse_rel_comp}
+\rulename{converse_relcomp}
 \end{isabelle}
 
 \indexbold{image!under a relation}%
--- a/src/Doc/manual.bib	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/Doc/manual.bib	Mon Nov 11 17:44:21 2013 +0100
@@ -926,7 +926,7 @@
   note = "\url{https://github.com/frelindb/agsyHOL}"}
 
 @incollection{lochbihler-2010,
-  title = "Coinduction",
+  title = "Coinductive",
   author = "Andreas Lochbihler",
   booktitle = "The Archive of Formal Proofs",
   editor = "Gerwin Klein and Tobias Nipkow and Lawrence C. Paulson",
--- a/src/HOL/ATP.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/ATP.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -18,34 +18,34 @@
 
 subsection {* Higher-order reasoning helpers *}
 
-definition fFalse :: bool where [no_atp]:
+definition fFalse :: bool where
 "fFalse \<longleftrightarrow> False"
 
-definition fTrue :: bool where [no_atp]:
+definition fTrue :: bool where
 "fTrue \<longleftrightarrow> True"
 
-definition fNot :: "bool \<Rightarrow> bool" where [no_atp]:
+definition fNot :: "bool \<Rightarrow> bool" where
 "fNot P \<longleftrightarrow> \<not> P"
 
-definition fComp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" where [no_atp]:
+definition fComp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" where
 "fComp P = (\<lambda>x. \<not> P x)"
 
-definition fconj :: "bool \<Rightarrow> bool \<Rightarrow> bool" where [no_atp]:
+definition fconj :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
 "fconj P Q \<longleftrightarrow> P \<and> Q"
 
-definition fdisj :: "bool \<Rightarrow> bool \<Rightarrow> bool" where [no_atp]:
+definition fdisj :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
 "fdisj P Q \<longleftrightarrow> P \<or> Q"
 
-definition fimplies :: "bool \<Rightarrow> bool \<Rightarrow> bool" where [no_atp]:
+definition fimplies :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
 "fimplies P Q \<longleftrightarrow> (P \<longrightarrow> Q)"
 
-definition fequal :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where [no_atp]:
+definition fequal :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
 "fequal x y \<longleftrightarrow> (x = y)"
 
-definition fAll :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where [no_atp]:
+definition fAll :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where
 "fAll P \<longleftrightarrow> All P"
 
-definition fEx :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where [no_atp]:
+definition fEx :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where
 "fEx P \<longleftrightarrow> Ex P"
 
 lemma fTrue_ne_fFalse: "fFalse \<noteq> fTrue"
--- a/src/HOL/Archimedean_Field.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Archimedean_Field.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -129,12 +129,8 @@
   fix y z assume
     "of_int y \<le> x \<and> x < of_int (y + 1)"
     "of_int z \<le> x \<and> x < of_int (z + 1)"
-  then have
-    "of_int y \<le> x" "x < of_int (y + 1)"
-    "of_int z \<le> x" "x < of_int (z + 1)"
-    by simp_all
-  from le_less_trans [OF `of_int y \<le> x` `x < of_int (z + 1)`]
-       le_less_trans [OF `of_int z \<le> x` `x < of_int (y + 1)`]
+  with le_less_trans [of "of_int y" "x" "of_int (z + 1)"]
+       le_less_trans [of "of_int z" "x" "of_int (y + 1)"]
   show "y = z" by (simp del: of_int_add)
 qed
 
--- a/src/HOL/BNF/BNF_FP_Base.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/BNF_FP_Base.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -172,7 +172,5 @@
 ML_file "Tools/bnf_fp_n2m.ML"
 ML_file "Tools/bnf_fp_n2m_sugar.ML"
 ML_file "Tools/bnf_fp_rec_sugar_util.ML"
-ML_file "Tools/bnf_fp_rec_sugar_tactics.ML"
-ML_file "Tools/bnf_fp_rec_sugar.ML"
 
 end
--- a/src/HOL/BNF/BNF_GFP.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/BNF_GFP.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -308,6 +308,8 @@
 lemma fun_rel_image2p: "(fun_rel R (image2p f g R)) f g"
   unfolding fun_rel_def image2p_def by auto
 
+ML_file "Tools/bnf_gfp_rec_sugar_tactics.ML"
+ML_file "Tools/bnf_gfp_rec_sugar.ML"
 ML_file "Tools/bnf_gfp_util.ML"
 ML_file "Tools/bnf_gfp_tactics.ML"
 ML_file "Tools/bnf_gfp.ML"
--- a/src/HOL/BNF/BNF_LFP.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/BNF_LFP.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -230,6 +230,7 @@
 lemma predicate2D_vimage2p: "\<lbrakk>R \<le> vimage2p f g S; R x y\<rbrakk> \<Longrightarrow> S (f x) (g y)"
   unfolding vimage2p_def by auto
 
+ML_file "Tools/bnf_lfp_rec_sugar.ML"
 ML_file "Tools/bnf_lfp_util.ML"
 ML_file "Tools/bnf_lfp_tactics.ML"
 ML_file "Tools/bnf_lfp.ML"
--- a/src/HOL/BNF/Basic_BNFs.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Basic_BNFs.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -27,15 +27,14 @@
 lemma wpull_Grp_def: "wpull A B1 B2 f1 f2 p1 p2 \<longleftrightarrow> Grp B1 f1 OO (Grp B2 f2)\<inverse>\<inverse> \<le> (Grp A p1)\<inverse>\<inverse> OO Grp A p2"
   unfolding wpull_def Grp_def by auto
 
-bnf ID: "id :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" ["\<lambda>x. {x}"] "\<lambda>_:: 'a. natLeq" ["id :: 'a \<Rightarrow> 'a"]
+bnf ID: "id :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" ["\<lambda>x. {x}"] "\<lambda>_:: 'a. natLeq"
   "id :: ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
 apply (auto simp: Grp_def fun_eq_iff relcompp.simps natLeq_card_order natLeq_cinfinite)
 apply (rule ordLess_imp_ordLeq[OF finite_ordLess_infinite[OF _ natLeq_Well_order]])
 apply (auto simp add: Field_card_of Field_natLeq card_of_well_order_on)[3]
 done
 
-bnf DEADID: "id :: 'a \<Rightarrow> 'a" [] "\<lambda>_:: 'a. natLeq +c |UNIV :: 'a set|" ["SOME x :: 'a. True"]
-  "op = :: 'a \<Rightarrow> 'a \<Rightarrow> bool"
+bnf DEADID: "id :: 'a \<Rightarrow> 'a" [] "\<lambda>_:: 'a. natLeq +c |UNIV :: 'a set|" "op = :: 'a \<Rightarrow> 'a \<Rightarrow> bool"
 by (auto simp add: wpull_Grp_def Grp_def
   card_order_csum natLeq_card_order card_of_card_order_on
   cinfinite_csum natLeq_cinfinite)
@@ -148,7 +147,7 @@
 
 lemmas prod_set_defs = fsts_def[abs_def] snds_def[abs_def]
 
-bnf map_pair [fsts, snds] "\<lambda>_::'a \<times> 'b. natLeq" [Pair] prod_rel
+bnf map_pair [fsts, snds] "\<lambda>_::'a \<times> 'b. natLeq" prod_rel
 proof (unfold prod_set_defs)
   show "map_pair id id = id" by (rule map_pair.id)
 next
@@ -193,7 +192,7 @@
         Grp {x. {fst x} \<subseteq> Collect (split R) \<and> {snd x} \<subseteq> Collect (split S)} (map_pair snd snd)"
   unfolding prod_set_defs prod_rel_def Grp_def relcompp.simps conversep.simps fun_eq_iff
   by auto
-qed simp+
+qed
 
 (* Categorical version of pullback: *)
 lemma wpull_cat:
@@ -231,7 +230,7 @@
   ultimately show ?thesis using card_of_ordLeq by fast
 qed
 
-bnf "op \<circ>" [range] "\<lambda>_:: 'a \<Rightarrow> 'b. natLeq +c |UNIV :: 'a set|" ["%c x::'b::type. c::'a::type"]
+bnf "op \<circ>" [range] "\<lambda>_:: 'a \<Rightarrow> 'b. natLeq +c |UNIV :: 'a set|"
   "fun_rel op ="
 proof
   fix f show "id \<circ> f = id f" by simp
@@ -278,6 +277,6 @@
          Grp {x. range x \<subseteq> Collect (split R)} (op \<circ> snd)"
   unfolding fun_rel_def Grp_def fun_eq_iff relcompp.simps conversep.simps  subset_iff image_iff
   by auto (force, metis pair_collapse)
-qed auto
+qed
 
 end
--- a/src/HOL/BNF/Examples/Misc_Codatatype.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Examples/Misc_Codatatype.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -19,9 +19,9 @@
 
 codatatype simple'' = X1'' nat int | X2''
 
-codatatype 'a stream = Stream 'a "'a stream"
+codatatype 'a stream = Stream (shd: 'a) (stl: "'a stream")
 
-codatatype 'a mylist = MyNil | MyCons 'a "'a mylist"
+codatatype 'a mylist = MyNil | MyCons (myhd: 'a) (mytl: "'a mylist")
 
 codatatype ('b, 'c, 'd, 'e) some_passive =
   SP1 "('b, 'c, 'd, 'e) some_passive" | SP2 'b | SP3 'c | SP4 'd | SP5 'e
--- a/src/HOL/BNF/Examples/Misc_Datatype.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Examples/Misc_Datatype.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -19,7 +19,7 @@
 
 datatype_new simple'' = X1'' nat int | X2''
 
-datatype_new 'a mylist = MyNil | MyCons 'a "'a mylist"
+datatype_new 'a mylist = MyNil | MyCons (myhd: 'a) (mytl: "'a mylist")
 
 datatype_new ('b, 'c, 'd, 'e) some_passive =
   SP1 "('b, 'c, 'd, 'e) some_passive" | SP2 'b | SP3 'c | SP4 'd | SP5 'e
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Examples/Misc_Primcorec.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -0,0 +1,112 @@
+(*  Title:      HOL/BNF/Examples/Misc_Primcorec.thy
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2013
+
+Miscellaneous primitive corecursive function definitions.
+*)
+
+header {* Miscellaneous Primitive Corecursive Function Definitions *}
+
+theory Misc_Primcorec
+imports Misc_Codatatype
+begin
+
+primcorec simple_of_bools :: "bool \<Rightarrow> bool \<Rightarrow> simple" where
+  "simple_of_bools b b' = (if b then if b' then X1 else X2 else if b' then X3 else X4)"
+
+primcorec simple'_of_bools :: "bool \<Rightarrow> bool \<Rightarrow> simple'" where
+  "simple'_of_bools b b' =
+     (if b then if b' then X1' () else X2' () else if b' then X3' () else X4' ())"
+
+primcorec inc_simple'' :: "nat \<Rightarrow> simple'' \<Rightarrow> simple''" where
+  "inc_simple'' k s = (case s of X1'' n i \<Rightarrow> X1'' (n + k) (i + int k) | X2'' \<Rightarrow> X2'')"
+
+primcorec sinterleave :: "'a stream \<Rightarrow> 'a stream \<Rightarrow> 'a stream" where
+  "sinterleave s s' = Stream (shd s) (sinterleave s' (stl s))"
+
+primcorec myapp :: "'a mylist \<Rightarrow> 'a mylist \<Rightarrow> 'a mylist" where
+  "myapp xs ys =
+     (if xs = MyNil then ys
+      else if ys = MyNil then xs
+      else MyCons (myhd xs) (myapp (mytl xs) ys))"
+
+primcorec shuffle_sp :: "('a, 'b, 'c, 'd) some_passive \<Rightarrow> ('d, 'a, 'b, 'c) some_passive" where
+  "shuffle_sp sp =
+     (case sp of
+       SP1 sp' \<Rightarrow> SP1 (shuffle_sp sp')
+     | SP2 a \<Rightarrow> SP3 a
+     | SP3 b \<Rightarrow> SP4 b
+     | SP4 c \<Rightarrow> SP5 c
+     | SP5 d \<Rightarrow> SP2 d)"
+
+primcorec rename_lam :: "(string \<Rightarrow> string) \<Rightarrow> lambda \<Rightarrow> lambda" where
+  "rename_lam f l =
+     (case l of
+       Var s \<Rightarrow> Var (f s)
+     | App l l' \<Rightarrow> App (rename_lam f l) (rename_lam f l')
+     | Abs s l \<Rightarrow> Abs (f s) (rename_lam f l)
+     | Let SL l \<Rightarrow> Let (fimage (map_pair f (rename_lam f)) SL) (rename_lam f l))"
+
+primcorec
+  j1_sum :: "('a\<Colon>{zero,one,plus}) \<Rightarrow> 'a J1" and
+  j2_sum :: "'a \<Rightarrow> 'a J2"
+where
+  "n = 0 \<Longrightarrow> is_J11 (j1_sum n)" |
+  "un_J111 (j1_sum _) = 0" |
+  "un_J112 (j1_sum _) = j1_sum 0" |
+  "un_J121 (j1_sum n) = n + 1" |
+  "un_J122 (j1_sum n) = j2_sum (n + 1)" |
+  "n = 0 \<Longrightarrow> is_J21 (j2_sum n)" |
+  "un_J221 (j2_sum n) = j1_sum (n + 1)" |
+  "un_J222 (j2_sum n) = j2_sum (n + 1)"
+
+primcorec forest_of_mylist :: "'a tree mylist \<Rightarrow> 'a forest" where
+  "forest_of_mylist ts =
+     (case ts of
+       MyNil \<Rightarrow> FNil
+     | MyCons t ts \<Rightarrow> FCons t (forest_of_mylist ts))"
+
+primcorec mylist_of_forest :: "'a forest \<Rightarrow> 'a tree mylist" where
+  "mylist_of_forest f =
+     (case f of
+       FNil \<Rightarrow> MyNil
+     | FCons t ts \<Rightarrow> MyCons t (mylist_of_forest ts))"
+
+primcorec semi_stream :: "'a stream \<Rightarrow> 'a stream" where
+  "semi_stream s = Stream (shd s) (semi_stream (stl (stl s)))"
+
+primcorec
+  tree'_of_stream :: "'a stream \<Rightarrow> 'a tree'" and
+  branch_of_stream :: "'a stream \<Rightarrow> 'a branch"
+where
+  "tree'_of_stream s =
+     TNode' (branch_of_stream (semi_stream s)) (branch_of_stream (semi_stream (stl s)))" |
+  "branch_of_stream s = (case s of Stream h t \<Rightarrow> Branch h (tree'_of_stream t))"
+
+primcorec
+  freeze_exp :: "('b \<Rightarrow> 'a) \<Rightarrow> ('a, 'b) exp \<Rightarrow> ('a, 'b) exp" and
+  freeze_trm :: "('b \<Rightarrow> 'a) \<Rightarrow> ('a, 'b) trm \<Rightarrow> ('a, 'b) trm" and
+  freeze_factor :: "('b \<Rightarrow> 'a) \<Rightarrow> ('a, 'b) factor \<Rightarrow> ('a, 'b) factor"
+where
+  "freeze_exp g e =
+     (case e of
+       Term t \<Rightarrow> Term (freeze_trm g t)
+     | Sum t e \<Rightarrow> Sum (freeze_trm g t) (freeze_exp g e))" |
+  "freeze_trm g t =
+     (case t of
+       Factor f \<Rightarrow> Factor (freeze_factor g f)
+     | Prod f t \<Rightarrow> Prod (freeze_factor g f) (freeze_trm g t))" |
+  "freeze_factor g f =
+     (case f of
+       C a \<Rightarrow> C a
+     | V b \<Rightarrow> C (g b)
+     | Paren e \<Rightarrow> Paren (freeze_exp g e))"
+
+primcorec poly_unity :: "'a poly_unit" where
+  "poly_unity = U (\<lambda>_. poly_unity)"
+
+primcorec build_cps :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> bool stream) \<Rightarrow> 'a \<Rightarrow> bool stream \<Rightarrow> 'a cps" where
+  "shd b \<Longrightarrow> build_cps f g a b = CPS1 a" |
+  "_ \<Longrightarrow> build_cps f g a b = CPS2 (\<lambda>a. build_cps f g (f a) (g a))"
+
+end
--- a/src/HOL/BNF/Examples/Misc_Primrec.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Examples/Misc_Primrec.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -14,7 +14,7 @@
 primrec_new nat_of_simple :: "simple \<Rightarrow> nat" where
   "nat_of_simple X1 = 1" |
   "nat_of_simple X2 = 2" |
-  "nat_of_simple X3 = 2" |
+  "nat_of_simple X3 = 3" |
   "nat_of_simple X4 = 4"
 
 primrec_new simple_of_simple' :: "simple' \<Rightarrow> simple" where
--- a/src/HOL/BNF/Tools/bnf_comp_tactics.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Tools/bnf_comp_tactics.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -164,10 +164,9 @@
 fun mk_comp_wit_tac ctxt Gwit_thms collect_set_map Fwit_thms =
   ALLGOALS (dtac @{thm in_Union_o_assoc}) THEN
   unfold_thms_tac ctxt (collect_set_map :: comp_wit_thms) THEN
-  REPEAT_DETERM (
-    atac 1 ORELSE
-    REPEAT_DETERM (eresolve_tac @{thms UnionE UnE imageE} 1) THEN
-    (TRY o dresolve_tac Gwit_thms THEN'
+  REPEAT_DETERM ((atac ORELSE'
+    REPEAT_DETERM o eresolve_tac @{thms UnionE UnE} THEN'
+    etac imageE THEN' TRY o dresolve_tac Gwit_thms THEN'
     (etac FalseE ORELSE'
     hyp_subst_tac ctxt THEN'
     dresolve_tac Fwit_thms THEN'
--- a/src/HOL/BNF/Tools/bnf_def.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Tools/bnf_def.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -77,14 +77,20 @@
   val wit_thms_of_bnf: bnf -> thm list
   val wit_thmss_of_bnf: bnf -> thm list list
 
+  val mk_map: int -> typ list -> typ list -> term -> term
+  val mk_rel: int -> typ list -> typ list -> term -> term
+  val build_map: Proof.context -> (typ * typ -> term) -> typ * typ -> term
+  val build_rel: Proof.context -> (typ * typ -> term) -> typ * typ -> term
+  val flatten_type_args_of_bnf: bnf -> 'a -> 'a list -> 'a list
+  val map_flattened_map_args: Proof.context -> string -> (term list -> 'a list) -> term list ->
+    'a list
+
   val mk_witness: int list * term -> thm list -> nonemptiness_witness
   val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list
   val wits_of_bnf: bnf -> nonemptiness_witness list
 
   val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list
 
-  val flatten_type_args_of_bnf: bnf -> 'a -> 'a list -> 'a list
-
   datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline
   datatype fact_policy = Dont_Note | Note_Some | Note_All
 
@@ -447,7 +453,6 @@
   #> Option.map (morph_bnf (Morphism.thm_morphism (Thm.transfer (Proof_Context.theory_of ctxt))));
 
 
-
 (* Utilities *)
 
 fun normalize_set insts instA set =
@@ -487,6 +492,46 @@
        else minimize ((I, wit) :: done) todo;
  in minimize [] wits end;
 
+fun mk_map live Ts Us t =
+  let val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last in
+    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
+  end;
+
+fun mk_rel live Ts Us t =
+  let val [Type (_, Ts0), Type (_, Us0)] = binder_types (snd (strip_typeN live (fastype_of t))) in
+    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
+  end;
+
+fun build_map_or_rel mk const of_bnf dest ctxt build_simple =
+  let
+    fun build (TU as (T, U)) =
+      if T = U then
+        const T
+      else
+        (case TU of
+          (Type (s, Ts), Type (s', Us)) =>
+          if s = s' then
+            let
+              val bnf = the (bnf_of ctxt s);
+              val live = live_of_bnf bnf;
+              val mapx = mk live Ts Us (of_bnf bnf);
+              val TUs' = map dest (fst (strip_typeN live (fastype_of mapx)));
+            in Term.list_comb (mapx, map build TUs') end
+          else
+            build_simple TU
+        | _ => build_simple TU);
+  in build end;
+
+val build_map = build_map_or_rel mk_map HOLogic.id_const map_of_bnf dest_funT;
+val build_rel = build_map_or_rel mk_rel HOLogic.eq_const rel_of_bnf dest_pred2T;
+
+fun map_flattened_map_args ctxt s map_args fs =
+  let
+    val flat_fs = flatten_type_args_of_bnf (the (bnf_of ctxt s)) Term.dummy fs;
+    val flat_fs' = map_args flat_fs;
+  in
+    permute_like (op aconv) flat_fs fs flat_fs'
+  end;
 
 
 (* Names *)
@@ -612,14 +657,12 @@
     val fact_policy = mk_fact_policy no_defs_lthy;
     val bnf_b = qualify raw_bnf_b;
     val live = length raw_sets;
-    val nwits = length raw_wits;
 
     val map_rhs = prep_term no_defs_lthy raw_map;
     val set_rhss = map (prep_term no_defs_lthy) raw_sets;
     val (bd_rhsT, bd_rhs) = (case prep_term no_defs_lthy raw_bd_Abs of
       Abs (_, T, t) => (T, t)
     | _ => error "Bad bound constant");
-    val wit_rhss = map (prep_term no_defs_lthy) raw_wits;
 
     fun err T =
       error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
@@ -633,7 +676,7 @@
         | T => err T)
       else (bnf_b, Local_Theory.full_name no_defs_lthy bnf_b);
 
-    val def_qualify = Binding.qualify false (Binding.name_of bnf_b);
+    val def_qualify = Binding.conceal o Binding.qualify false (Binding.name_of bnf_b);
 
     fun mk_suffix_binding suf = Binding.suffix_name ("_" ^ suf) bnf_b;
 
@@ -672,21 +715,14 @@
           else map (fn i => set_name i (fn () => mk_suffix_binding (mk_setN i))) (1 upto live);
       in bs ~~ set_rhss end;
     val bd_bind_def = (fn () => def_qualify (mk_suffix_binding bdN), bd_rhs);
-    val wit_binds_defs =
-      let
-        val bs = if nwits = 1 then [fn () => def_qualify (mk_suffix_binding witN)]
-          else map (fn i => fn () => def_qualify (mk_suffix_binding (mk_witN i))) (1 upto nwits);
-      in bs ~~ wit_rhss end;
 
-    val (((((bnf_map_term, raw_map_def),
+    val ((((bnf_map_term, raw_map_def),
       (bnf_set_terms, raw_set_defs)),
-      (bnf_bd_term, raw_bd_def)),
-      (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) =
+      (bnf_bd_term, raw_bd_def)), (lthy, lthy_old)) =
         no_defs_lthy
         |> maybe_define true map_bind_def
         ||>> apfst split_list o fold_map (maybe_define true) set_binds_defs
         ||>> maybe_define true bd_bind_def
-        ||>> apfst split_list o fold_map (maybe_define true) wit_binds_defs
         ||> `(maybe_restore no_defs_lthy);
 
     val phi = Proof_Context.export_morphism lthy_old lthy;
@@ -694,7 +730,6 @@
     val bnf_map_def = Morphism.thm phi raw_map_def;
     val bnf_set_defs = map (Morphism.thm phi) raw_set_defs;
     val bnf_bd_def = Morphism.thm phi raw_bd_def;
-    val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs;
 
     val bnf_map = Morphism.term phi bnf_map_term;
 
@@ -713,7 +748,6 @@
     val bdT = Morphism.typ phi bd_rhsT;
     val bnf_bd =
       Term.subst_TVars (Term.add_tvar_namesT bdT [] ~~ CA_params) (Morphism.term phi bnf_bd_term);
-    val bnf_wits = map (normalize_wit CA_params CA alphas o Morphism.term phi) bnf_wit_terms;
 
     (*TODO: assert Ds = (TVars of bnf_map) \ (alphas @ betas) as sets*)
     val deads = (case Ds_opt of
@@ -770,7 +804,6 @@
     val bnf_sets_As = map (mk_bnf_t As') bnf_sets;
     val bnf_sets_Bs = map (mk_bnf_t Bs') bnf_sets;
     val bnf_bd_As = mk_bnf_t As' bnf_bd;
-    val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits;
 
     val pre_names_lthy = lthy;
     val ((((((((((((((((((((((((fs, gs), hs), x), y), zs), ys), As),
@@ -827,9 +860,23 @@
       (fn () => def_qualify (if Binding.is_empty rel_b then mk_suffix_binding relN else rel_b),
          rel_rhs);
 
-    val ((bnf_rel_term, raw_rel_def), (lthy, lthy_old)) =
+    val wit_rhss =
+      if null raw_wits then
+        [fold_rev Term.absdummy As' (Term.list_comb (bnf_map_AsAs,
+          map2 (fn T => fn i => Term.absdummy T (Bound i)) As' (live downto 1)) $
+          Const (@{const_name undefined}, CA'))]
+      else map (prep_term no_defs_lthy) raw_wits;
+    val nwits = length wit_rhss;
+    val wit_binds_defs =
+      let
+        val bs = if nwits = 1 then [fn () => def_qualify (mk_suffix_binding witN)]
+          else map (fn i => fn () => def_qualify (mk_suffix_binding (mk_witN i))) (1 upto nwits);
+      in bs ~~ wit_rhss end;
+
+    val (((bnf_rel_term, raw_rel_def), (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) =
       lthy
       |> maybe_define (is_some raw_rel_opt) rel_bind_def
+      ||>> apfst split_list o fold_map (maybe_define (not (null raw_wits))) wit_binds_defs
       ||> `(maybe_restore lthy);
 
     val phi = Proof_Context.export_morphism lthy_old lthy;
@@ -841,11 +888,9 @@
     val rel = mk_bnf_rel pred2RTs CA' CB';
     val relAsAs = mk_bnf_rel self_pred2RTs CA' CA';
 
-    val _ = case no_reflexive (raw_map_def :: raw_set_defs @ [raw_bd_def] @
-        raw_wit_defs @ [raw_rel_def]) of
-        [] => ()
-      | defs => Proof_Display.print_consts true lthy_old (K false)
-          (map (dest_Free o fst o Logic.dest_equals o prop_of) defs);
+    val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs;
+    val bnf_wits = map (normalize_wit CA_params CA alphas o Morphism.term phi) bnf_wit_terms;
+    val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits;
 
     val map_id0_goal =
       let val bnf_map_app_id = Term.list_comb (bnf_map_AsAs, map HOLogic.id_const As') in
@@ -945,11 +990,14 @@
         map wit_goal (0 upto live - 1)
       end;
 
-    val wit_goalss = map mk_wit_goals bnf_wit_As;
+    val trivial_wit_tac = mk_trivial_wit_tac bnf_wit_defs;
 
-    fun after_qed thms lthy =
+    val wit_goalss =
+      (if null raw_wits then SOME trivial_wit_tac else NONE, map mk_wit_goals bnf_wit_As);
+
+    fun after_qed mk_wit_thms thms lthy =
       let
-        val (axioms, wit_thms) = apfst (mk_axioms live) (chop (length goals) thms);
+        val (axioms, nontriv_wit_thms) = apfst (mk_axioms live) (chop (length goals) thms);
 
         val bd_Card_order = #bd_card_order axioms RS @{thm conjunct2[OF card_order_on_Card_order]};
         val bd_Cinfinite = @{thm conjI} OF [#bd_cinfinite axioms, bd_Card_order];
@@ -1022,6 +1070,9 @@
 
         val set_map = map (fn thm => Lazy.lazy (fn () => mk_set_map thm)) (#set_map0 axioms);
 
+        val wit_thms =
+          if null nontriv_wit_thms then mk_wit_thms (map Lazy.force set_map) else nontriv_wit_thms;
+
         fun mk_in_bd () =
           let
             val bdT = fst (dest_relT (fastype_of bnf_bd_As));
@@ -1265,35 +1316,45 @@
   (bnf, Local_Theory.declaration {syntax = false, pervasive = true}
     (fn phi => Data.map (Symtab.default (key, morph_bnf phi bnf))) lthy);
 
-(* TODO: Once the invariant "nwits > 0" holds, remove "mk_conjunction_balanced'" and "rtac TrueI"
-   below *)
-fun mk_conjunction_balanced' [] = @{prop True}
-  | mk_conjunction_balanced' ts = Logic.mk_conjunction_balanced ts;
-
 fun bnf_def const_policy fact_policy qualify tacs wit_tac Ds map_b rel_b set_bs =
-  (fn (_, goals, wit_goalss, after_qed, lthy, one_step_defs) =>
+  (fn (_, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, one_step_defs) =>
   let
-    val wits_tac =
-      K (TRYALL Goal.conjunction_tac) THEN' K (TRYALL (rtac TrueI)) THEN'
-      mk_unfold_thms_then_tac lthy one_step_defs wit_tac;
-    val wit_goals = map mk_conjunction_balanced' wit_goalss;
-    val wit_thms =
-      Goal.prove_sorry lthy [] [] (mk_conjunction_balanced' wit_goals) wits_tac
-      |> Conjunction.elim_balanced (length wit_goals)
-      |> map2 (Conjunction.elim_balanced o length) wit_goalss
-      |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
+    fun mk_wits_tac set_maps =
+      K (TRYALL Goal.conjunction_tac) THEN'
+      (case triv_tac_opt of
+        SOME tac => tac set_maps
+      | NONE => mk_unfold_thms_then_tac lthy one_step_defs wit_tac);
+    val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;
+    fun mk_wit_thms set_maps =
+      Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals) (mk_wits_tac set_maps)
+        |> Conjunction.elim_balanced (length wit_goals)
+        |> map2 (Conjunction.elim_balanced o length) wit_goalss
+        |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
   in
     map2 (Thm.close_derivation oo Goal.prove_sorry lthy [] [])
       goals (map (mk_unfold_thms_then_tac lthy one_step_defs) tacs)
-    |> (fn thms => after_qed (map single thms @ wit_thms) lthy)
+    |> (fn thms => after_qed mk_wit_thms (map single thms) lthy)
   end) oo prepare_def const_policy fact_policy qualify (K I) Ds map_b rel_b set_bs;
 
-val bnf_cmd = (fn (key, goals, wit_goals, after_qed, lthy, defs) =>
-  Proof.unfolding ([[(defs, [])]])
-    (Proof.theorem NONE (snd o register_bnf key oo after_qed)
-      (map (single o rpair []) goals @ map (map (rpair [])) wit_goals) lthy)) oo
-  prepare_def Do_Inline (user_policy Note_Some) I Syntax.read_term NONE Binding.empty Binding.empty
-    [];
+val bnf_cmd = (fn (key, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, defs) =>
+  let
+    val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;
+    fun mk_triv_wit_thms tac set_maps =
+      Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals)
+        (K (TRYALL Goal.conjunction_tac) THEN' tac set_maps)
+        |> Conjunction.elim_balanced (length wit_goals)
+        |> map2 (Conjunction.elim_balanced o length) wit_goalss
+        |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
+    val (mk_wit_thms, nontriv_wit_goals) = 
+      (case triv_tac_opt of
+        NONE => (fn _ => [], map (map (rpair [])) wit_goalss)
+      | SOME tac => (mk_triv_wit_thms tac, []));
+  in
+    Proof.unfolding ([[(defs, [])]])
+      (Proof.theorem NONE (snd o register_bnf key oo after_qed mk_wit_thms)
+        (map (single o rpair []) goals @ nontriv_wit_goals) lthy)
+  end) oo prepare_def Do_Inline (user_policy Note_Some) I Syntax.read_term NONE
+    Binding.empty Binding.empty [];
 
 fun print_bnfs ctxt =
   let
@@ -1330,7 +1391,9 @@
     "register a type as a bounded natural functor"
     ((parse_opt_binding_colon -- Parse.term --
        (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) -- Parse.term --
-       (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) -- Scan.option Parse.term)
+       (Scan.option ((@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}))
+         >> the_default []) --
+       Scan.option Parse.term)
        >> bnf_cmd);
 
 end;
--- a/src/HOL/BNF/Tools/bnf_def_tactics.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Tools/bnf_def_tactics.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -31,7 +31,10 @@
     {prems: thm list, context: Proof.context} -> tactic
 
   val mk_in_bd_tac: int -> thm -> thm -> thm -> thm -> thm list -> thm list -> thm -> thm -> thm ->
-    thm -> {prems: 'a, context: Proof.context} -> tactic
+    thm -> {prems: thm list, context: Proof.context} -> tactic
+
+  val mk_trivial_wit_tac: thm list -> thm list -> {prems: thm list, context: Proof.context} ->
+    tactic
 end;
 
 structure BNF_Def_Tactics : BNF_DEF_TACTICS =
@@ -302,4 +305,8 @@
            map_comp RS sym, map_id])] 1
   end;
 
+fun mk_trivial_wit_tac wit_defs set_maps {context = ctxt, prems = _} =
+  unfold_thms_tac ctxt wit_defs THEN HEADGOAL (EVERY' (map (fn thm =>
+    dtac (thm RS equalityD1 RS set_mp) THEN' etac imageE THEN' atac) set_maps)) THEN ALLGOALS atac;
+
 end;
--- a/src/HOL/BNF/Tools/bnf_fp_def_sugar.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Tools/bnf_fp_def_sugar.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -25,7 +25,9 @@
      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
   val fp_sugars_of: Proof.context -> fp_sugar list
 
@@ -39,17 +41,14 @@
     'a list
   val mk_co_iter: theory -> BNF_FP_Util.fp_kind -> typ -> typ list -> term -> term
   val nesty_bnfs: Proof.context -> typ list list list -> typ list -> BNF_Def.bnf list
-  val mk_map: int -> typ list -> typ list -> term -> term
-  val mk_rel: int -> typ list -> typ list -> term -> term
-  val build_map: local_theory -> (typ * typ -> term) -> typ * typ -> term
-  val build_rel: local_theory -> (typ * typ -> term) -> typ * typ -> term
-  val dest_map: Proof.context -> string -> term -> term * term list
-  val dest_ctr: Proof.context -> string -> term -> term * term list
 
   type lfp_sugar_thms =
     (thm list * thm * Args.src list)
     * (thm list list * thm list list * Args.src list)
 
+  val morph_lfp_sugar_thms: morphism -> lfp_sugar_thms -> lfp_sugar_thms
+  val transfer_lfp_sugar_thms: Proof.context -> lfp_sugar_thms -> lfp_sugar_thms
+
   type gfp_sugar_thms =
     ((thm list * thm) list * Args.src list)
     * (thm list list * thm list list * Args.src list)
@@ -57,6 +56,9 @@
     * (thm list list * thm list list * Args.src list)
     * (thm list list list * thm list list list * Args.src list)
 
+  val morph_gfp_sugar_thms: morphism -> gfp_sugar_thms -> gfp_sugar_thms
+  val transfer_gfp_sugar_thms: Proof.context -> gfp_sugar_thms -> gfp_sugar_thms
+
   val mk_co_iters_prelims: BNF_FP_Util.fp_kind -> typ list list list -> typ list -> typ list ->
     int list -> int list list -> term list list -> Proof.context ->
     (term list list
@@ -87,8 +89,9 @@
     string * term list * term list list * ((term list list * term list list list)
       * (typ list * typ list list)) list ->
     thm -> thm list -> thm list -> thm list list -> BNF_Def.bnf list -> typ list -> typ list ->
-    int list list -> int list list -> int list -> thm list list -> Ctr_Sugar.ctr_sugar list ->
-    term list list -> thm list list -> (thm list -> thm list) -> local_theory -> gfp_sugar_thms
+    typ list -> typ list list list -> int list list -> int list list -> int list -> thm list list ->
+    Ctr_Sugar.ctr_sugar list -> term list list -> thm list list -> (thm list -> thm list) ->
+    local_theory -> gfp_sugar_thms
   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) ->
@@ -207,8 +210,8 @@
 val id_def = @{thm id_def};
 val mp_conj = @{thm mp_conj};
 
-val nitpick_attrs = @{attributes [nitpick_simp]};
-val code_nitpick_simp_attrs = Code.add_default_eqn_attrib :: nitpick_attrs;
+val nitpicksimp_attrs = @{attributes [nitpick_simp]};
+val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
 val simp_attrs = @{attributes [simp]};
 
 fun tvar_subst thy Ts Us =
@@ -232,7 +235,9 @@
   | flat_corec_preds_predsss_gettersss (p :: ps) (qss :: qsss) (fss :: fsss) =
     p :: flat_corec_predss_getterss qss fss @ flat_corec_preds_predsss_gettersss ps qsss fsss;
 
-fun mk_tupled_fun x f xs = HOLogic.tupled_lambda x (Term.list_comb (f, xs));
+fun mk_tupled_fun x f xs =
+  if xs = [x] then f else HOLogic.tupled_lambda x (Term.list_comb (f, xs));
+
 fun mk_uncurried2_fun f xss =
   mk_tupled_fun (HOLogic.mk_tuple (map HOLogic.mk_tuple xss)) f (flat_rec_arg_args xss);
 
@@ -287,66 +292,6 @@
   | unzip_corecT _ (Type (@{type_name sum}, Ts)) = Ts
   | unzip_corecT _ T = [T];
 
-fun mk_map live Ts Us t =
-  let val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last in
-    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
-  end;
-
-fun mk_rel live Ts Us t =
-  let val [Type (_, Ts0), Type (_, Us0)] = binder_types (snd (strip_typeN live (fastype_of t))) in
-    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
-  end;
-
-local
-
-fun build_map_or_rel mk const of_bnf dest lthy build_simple =
-  let
-    fun build (TU as (T, U)) =
-      if T = U then
-        const T
-      else
-        (case TU of
-          (Type (s, Ts), Type (s', Us)) =>
-          if s = s' then
-            let
-              val bnf = the (bnf_of lthy s);
-              val live = live_of_bnf bnf;
-              val mapx = mk live Ts Us (of_bnf bnf);
-              val TUs' = map dest (fst (strip_typeN live (fastype_of mapx)));
-            in Term.list_comb (mapx, map build TUs') end
-          else
-            build_simple TU
-        | _ => build_simple TU);
-  in build end;
-
-in
-
-val build_map = build_map_or_rel mk_map HOLogic.id_const map_of_bnf dest_funT;
-val build_rel = build_map_or_rel mk_rel HOLogic.eq_const rel_of_bnf dest_pred2T;
-
-end;
-
-val dummy_var_name = "?f"
-
-fun mk_map_pattern ctxt s =
-  let
-    val bnf = the (bnf_of ctxt s);
-    val mapx = map_of_bnf bnf;
-    val live = live_of_bnf bnf;
-    val (f_Ts, _) = strip_typeN live (fastype_of mapx);
-    val fs = map_index (fn (i, T) => Var ((dummy_var_name, i), T)) f_Ts;
-  in
-    (mapx, betapplys (mapx, fs))
-  end;
-
-fun dest_map ctxt s call =
-  let
-    val (map0, pat) = mk_map_pattern ctxt s;
-    val (_, tenv) = fo_match ctxt call pat;
-  in
-    (map0, Vartab.fold_rev (fn (_, (_, f)) => cons f) tenv [])
-  end;
-
 fun liveness_of_fp_bnf n bnf =
   (case T_of_bnf bnf of
     Type (_, Ts) => map (not o member (op =) (deads_of_bnf bnf)) Ts
@@ -388,12 +333,19 @@
 fun nesty_bnfs ctxt ctr_Tsss Us =
   map_filter (bnf_of ctxt) (fold (fold (fold (add_nesty_bnf_names Us))) ctr_Tsss []);
 
-fun indexify proj xs f p = f (find_index (curry op = (proj p)) xs) p;
+fun indexify proj xs f p = f (find_index (curry (op =) (proj p)) xs) p;
 
 type lfp_sugar_thms =
   (thm list * thm * Args.src list)
   * (thm list list * thm list list * Args.src list)
 
+fun morph_lfp_sugar_thms phi ((inducts, induct, induct_attrs), (foldss, recss, iter_attrs)) =
+  ((map (Morphism.thm phi) inducts, Morphism.thm phi induct, induct_attrs),
+   (map (map (Morphism.thm phi)) foldss, map (map (Morphism.thm phi)) recss, iter_attrs));
+
+val transfer_lfp_sugar_thms =
+  morph_lfp_sugar_thms o Morphism.thm_morphism o Thm.transfer o Proof_Context.theory_of;
+
 type gfp_sugar_thms =
   ((thm list * thm) list * Args.src list)
   * (thm list list * thm list list * Args.src list)
@@ -401,6 +353,23 @@
   * (thm list list * thm list list * Args.src list)
   * (thm list list list * thm list list list * Args.src list);
 
+fun morph_gfp_sugar_thms phi ((coinducts_pairs, coinduct_attrs),
+    (unfoldss, corecss, coiter_attrs), (disc_unfoldss, disc_corecss, disc_iter_attrs),
+    (disc_unfold_iffss, disc_corec_iffss, disc_iter_iff_attrs),
+    (sel_unfoldsss, sel_corecsss, sel_iter_attrs)) =
+  ((map (apfst (map (Morphism.thm phi)) o apsnd (Morphism.thm phi)) coinducts_pairs,
+    coinduct_attrs),
+   (map (map (Morphism.thm phi)) unfoldss, map (map (Morphism.thm phi)) corecss, coiter_attrs),
+   (map (map (Morphism.thm phi)) disc_unfoldss, map (map (Morphism.thm phi)) disc_corecss,
+    disc_iter_attrs),
+   (map (map (Morphism.thm phi)) disc_unfold_iffss, map (map (Morphism.thm phi)) disc_corec_iffss,
+    disc_iter_iff_attrs),
+   (map (map (map (Morphism.thm phi))) sel_unfoldsss,
+    map (map (map (Morphism.thm phi))) sel_corecsss, sel_iter_attrs));
+
+val transfer_gfp_sugar_thms =
+  morph_gfp_sugar_thms o Morphism.thm_morphism o Thm.transfer o Proof_Context.theory_of;
+
 fun mk_iter_fun_arg_types0 n ms = map2 dest_tupleT ms o dest_sumTN_balanced n o domain_type;
 
 fun mk_iter_fun_arg_types ctr_Tsss ns mss =
@@ -430,7 +399,7 @@
         ns mss ctr_Tsss ctor_iter_fun_Tss;
 
     val z_Tsss' = map (map flat_rec_arg_args) z_Tssss;
-    val h_Tss = map2 (map2 (curry op --->)) z_Tsss' Css;
+    val h_Tss = map2 (map2 (curry (op --->))) z_Tsss' Css;
 
     val hss = map2 (map2 retype_free) h_Tss gss;
     val zssss_hd = map2 (map2 (map2 (retype_free o hd))) z_Tssss ysss;
@@ -452,7 +421,7 @@
     val f_sum_prod_Ts = map range_type fun_Ts;
     val f_prod_Tss = map2 dest_sumTN_balanced ns f_sum_prod_Ts;
     val f_Tsss = map2 (map2 (dest_tupleT o length)) ctr_Tsss' f_prod_Tss;
-    val f_Tssss = map3 (fn C => map2 (map2 (map (curry op --> C) oo unzip_corecT)))
+    val f_Tssss = map3 (fn C => map2 (map2 (map (curry (op -->) C) oo unzip_corecT)))
       Cs ctr_Tsss' f_Tsss;
     val q_Tssss = map (map (map (fn [_] => [] | [_, T] => [mk_pred1T (domain_type T)]))) f_Tssss;
   in
@@ -577,7 +546,7 @@
 
     fun generate_iter suf (f_Tss, _, fss, xssss) ctor_iter =
       let
-        val res_T = fold_rev (curry op --->) f_Tss fpT_to_C;
+        val res_T = fold_rev (curry (op --->)) f_Tss fpT_to_C;
         val b = mk_binding suf;
         val spec =
           mk_Trueprop_eq (lists_bmoc fss (Free (Binding.name_of b, res_T)),
@@ -596,7 +565,7 @@
 
     fun generate_coiter suf ((pfss, cqfsss), (f_sum_prod_Ts, pf_Tss)) dtor_coiter =
       let
-        val res_T = fold_rev (curry op --->) pf_Tss C_to_fpT;
+        val res_T = fold_rev (curry (op --->)) pf_Tss C_to_fpT;
         val b = mk_binding suf;
         val spec =
           mk_Trueprop_eq (lists_bmoc pfss (Free (Binding.name_of b, res_T)),
@@ -645,7 +614,7 @@
         val lives = lives_of_bnf bnf;
         val sets = sets_of_bnf bnf;
         fun mk_set U =
-          (case find_index (curry op = U) lives of
+          (case find_index (curry (op =) U) lives of
             ~1 => Term.dummy
           | i => nth sets i);
       in
@@ -662,7 +631,7 @@
           end;
 
         fun mk_raw_prem_prems _ (x as Free (_, Type _)) (X as TFree _) =
-            [([], (find_index (curry op = X) Xs + 1, x))]
+            [([], (find_index (curry (op =) X) Xs + 1, x))]
           | mk_raw_prem_prems names_lthy (x as Free (s, Type (T_name, Ts0))) (Type (_, Xs_Ts0)) =
             (case AList.lookup (op =) setss_nested T_name of
               NONE => []
@@ -702,7 +671,7 @@
 
         val goal =
           Library.foldr (Logic.list_implies o apfst (map mk_prem)) (raw_premss,
-            HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry op $) ps us)));
+            HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry (op $)) ps us)));
 
         val kksss = map (map (map (fst o snd) o #2)) raw_premss;
 
@@ -763,13 +732,13 @@
     val rec_thmss = mk_iter_thmss rec_args_types recs rec_defs (map co_rec_of ctor_iter_thmss);
   in
     ((induct_thms, induct_thm, [induct_case_names_attr]),
-     (fold_thmss, rec_thmss, code_nitpick_simp_attrs @ simp_attrs))
+     (fold_thmss, rec_thmss, code_nitpicksimp_attrs @ simp_attrs))
   end;
 
 fun derive_coinduct_coiters_thms_for_types pre_bnfs (z, cs, cpss,
       coiters_args_types as [((pgss, crgsss), _), ((phss, cshsss), _)])
-    dtor_coinduct dtor_injects dtor_ctors dtor_coiter_thmss nesting_bnfs fpTs Cs kss mss ns
-    ctr_defss (ctr_sugars : ctr_sugar list) coiterss coiter_defss export_args lthy =
+    dtor_coinduct dtor_injects dtor_ctors dtor_coiter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss
+    mss ns ctr_defss (ctr_sugars : ctr_sugar list) coiterss coiter_defss export_args lthy =
   let
     fun mk_ctor_dtor_coiter_thm dtor_inject dtor_ctor coiter =
       iffD1 OF [dtor_inject, trans OF [coiter, dtor_ctor RS sym]];
@@ -821,40 +790,29 @@
           map4 (fn u => fn v => fn uvr => fn uv_eq =>
             fold_rev Term.lambda [u, v] (HOLogic.mk_disj (uvr, uv_eq))) us vs uvrs uv_eqs;
 
-        (* TODO: generalize (cf. "build_map") *)
-        fun build_rel rs' T =
-          (case find_index (curry op = T) fpTs of
-            ~1 =>
-            if exists_subtype_in fpTs T then
-              let
-                val Type (s, Ts) = T
-                val bnf = the (bnf_of lthy s);
-                val live = live_of_bnf bnf;
-                val rel = mk_rel live Ts Ts (rel_of_bnf bnf);
-                val Ts' = map domain_type (fst (strip_typeN live (fastype_of rel)));
-              in Term.list_comb (rel, map (build_rel rs') Ts') end
-            else
-              HOLogic.eq_const T
-          | kk => nth rs' kk);
+        fun build_the_rel rs' T Xs_T =
+          build_rel lthy (fn (_, X) => nth rs' (find_index (curry (op =) X) Xs)) (T, Xs_T)
+          |> Term.subst_atomic_types (Xs ~~ fpTs);
 
-        fun build_rel_app rs' usel vsel = fold rapp [usel, vsel] (build_rel rs' (fastype_of usel));
+        fun build_rel_app rs' usel vsel Xs_T =
+          fold rapp [usel, vsel] (build_the_rel rs' (fastype_of usel) Xs_T);
 
-        fun mk_prem_ctr_concls rs' n k udisc usels vdisc vsels =
+        fun mk_prem_ctr_concls rs' n k udisc usels vdisc vsels ctrXs_Ts =
           (if k = n then [] else [HOLogic.mk_eq (udisc, vdisc)]) @
           (if null usels then
              []
            else
              [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [udisc, vdisc],
-                Library.foldr1 HOLogic.mk_conj (map2 (build_rel_app rs') usels vsels))]);
+                Library.foldr1 HOLogic.mk_conj (map3 (build_rel_app rs') usels vsels ctrXs_Ts))]);
 
-        fun mk_prem_concl rs' n udiscs uselss vdiscs vselss =
-          Library.foldr1 HOLogic.mk_conj
-            (flat (map5 (mk_prem_ctr_concls rs' n) (1 upto n) udiscs uselss vdiscs vselss))
+        fun mk_prem_concl rs' n udiscs uselss vdiscs vselss ctrXs_Tss =
+          Library.foldr1 HOLogic.mk_conj (flat (map6 (mk_prem_ctr_concls rs' n)
+            (1 upto n) udiscs uselss vdiscs vselss ctrXs_Tss))
           handle List.Empty => @{term True};
 
-        fun mk_prem rs' uvr u v n udiscs uselss vdiscs vselss =
+        fun mk_prem rs' uvr u v n udiscs uselss vdiscs vselss ctrXs_Tss =
           fold_rev Logic.all [u, v] (Logic.mk_implies (HOLogic.mk_Trueprop uvr,
-            HOLogic.mk_Trueprop (mk_prem_concl rs' n udiscs uselss vdiscs vselss)));
+            HOLogic.mk_Trueprop (mk_prem_concl rs' n udiscs uselss vdiscs vselss ctrXs_Tss)));
 
         val concl =
           HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
@@ -862,8 +820,8 @@
                uvrs us vs));
 
         fun mk_goal rs' =
-          Logic.list_implies (map8 (mk_prem rs') uvrs us vs ns udiscss uselsss vdiscss vselsss,
-            concl);
+          Logic.list_implies (map9 (mk_prem rs') uvrs us vs ns udiscss uselsss vdiscss vselsss
+            ctrXs_Tsss, concl);
 
         val goals = map mk_goal [rs, strong_rs];
 
@@ -1024,7 +982,7 @@
       coinduct_consumes_attr :: coinduct_case_names_attr :: coinduct_case_concl_attrs;
   in
     ((coinduct_thms_pairs, coinduct_case_attrs),
-     (unfold_thmss, corec_thmss, code_nitpick_simp_attrs),
+     (unfold_thmss, corec_thmss, code_nitpicksimp_attrs),
      (disc_unfold_thmss, disc_corec_thmss, []),
      (disc_unfold_iff_thmss, disc_corec_iff_thmss, simp_attrs),
      (sel_unfold_thmsss, sel_corec_thmsss, simp_attrs))
@@ -1074,7 +1032,7 @@
 
     val qsoty = quote o Syntax.string_of_typ fake_lthy;
 
-    val _ = (case duplicates (op =) unsorted_As of [] => ()
+    val _ = (case Library.duplicates (op =) unsorted_As of [] => ()
       | A :: _ => error ("Duplicate type parameter " ^ qsoty A ^ " in " ^ co_prefix fp ^
           "datatype specification"));
 
@@ -1087,7 +1045,7 @@
 
     val mixfixes = map mixfix_of specs;
 
-    val _ = (case duplicates Binding.eq_name fp_bs of [] => ()
+    val _ = (case Library.duplicates Binding.eq_name fp_bs of [] => ()
       | b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b)));
 
     val ctr_specss = map ctr_specs_of specs;
@@ -1380,15 +1338,22 @@
               val (rel_distinct_thms, _) =
                 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,
+                  code_nitpicksimp_attrs),
+                 (map2 (fn th => fn 0 => th RS @{thm eq_True[THEN iffD2]} | _ => th)
+                    rel_inject_thms ms, code_nitpicksimp_attrs)]
+                |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
+
               val notes =
-                [(mapN, map_thms, code_nitpick_simp_attrs @ simp_attrs),
-                 (rel_distinctN, rel_distinct_thms, code_nitpick_simp_attrs @ simp_attrs),
-                 (rel_injectN, rel_inject_thms, code_nitpick_simp_attrs @ simp_attrs),
-                 (setN, flat set_thmss, code_nitpick_simp_attrs @ simp_attrs)]
+                [(mapN, map_thms, code_nitpicksimp_attrs @ simp_attrs),
+                 (rel_distinctN, rel_distinct_thms, simp_attrs),
+                 (rel_injectN, rel_inject_thms, simp_attrs),
+                 (setN, flat set_thmss, code_nitpicksimp_attrs @ simp_attrs)]
                 |> massage_simple_notes fp_b_name;
             in
               (((map_thms, rel_inject_thms, rel_distinct_thms, set_thmss), ctr_sugar),
-               lthy |> Local_Theory.notes notes |> snd)
+               lthy |> Local_Theory.notes (anonymous_notes @ notes) |> snd)
             end;
 
         fun mk_binding suf = qualify false fp_b_name (Binding.suffix_name ("_" ^ suf) fp_b);
@@ -1457,8 +1422,9 @@
              (disc_unfold_iff_thmss, disc_corec_iff_thmss, disc_coiter_iff_attrs),
              (sel_unfold_thmsss, sel_corec_thmsss, sel_coiter_attrs)) =
           derive_coinduct_coiters_thms_for_types pre_bnfs (the coiters_args_types) xtor_co_induct
-            dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs kss mss ns ctr_defss
-            ctr_sugars coiterss coiter_defss (Proof_Context.export lthy' no_defs_lthy) lthy;
+            dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns
+            ctr_defss ctr_sugars coiterss coiter_defss (Proof_Context.export lthy' no_defs_lthy)
+            lthy;
 
         val sel_unfold_thmss = map flat sel_unfold_thmsss;
         val sel_corec_thmss = map flat sel_corec_thmsss;
@@ -1496,6 +1462,12 @@
            (unfoldN, unfold_thmss, K coiter_attrs)]
           |> massage_multi_notes;
 
+        fun is_codatatype (Type (s, _)) =
+            (case fp_sugar_of lthy s of SOME {fp = Greatest_FP, ...} => true | _ => false)
+          | is_codatatype _ = false;
+
+        val nitpick_supported = forall (is_codatatype o T_of_bnf) nested_bnfs;
+
         fun register_nitpick fpT ({ctrs, casex, ...} : ctr_sugar) =
           Nitpick_HOL.register_codatatype fpT (fst (dest_Const casex))
             (map (dest_Const o mk_ctr As) ctrs)
@@ -1507,7 +1479,7 @@
           ctr_sugars coiterss mapss [coinduct_thm, strong_coinduct_thm]
           (transpose [unfold_thmss, corec_thmss]) (transpose [disc_unfold_thmss, disc_corec_thmss])
           (transpose [sel_unfold_thmsss, sel_corec_thmsss])
-        |> fold2 register_nitpick fpTs ctr_sugars
+        |> nitpick_supported ? fold2 register_nitpick fpTs ctr_sugars
       end;
 
     val lthy'' = lthy'
--- a/src/HOL/BNF/Tools/bnf_fp_def_sugar_tactics.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Tools/bnf_fp_def_sugar_tactics.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -151,12 +151,17 @@
   (atac ORELSE' REPEAT o etac conjE THEN'
      full_simp_tac
        (ss_only (@{thm prod.inject} :: no_refl discs @ rel_eqs @ more_simp_thms) ctxt) THEN'
-     REPEAT o etac conjE THEN_MAYBE' REPEAT o hyp_subst_tac ctxt THEN' REPEAT o rtac conjI THEN'
-     REPEAT o (rtac refl ORELSE' atac));
+     REPEAT o etac conjE THEN_MAYBE' REPEAT o hyp_subst_tac ctxt THEN'
+     REPEAT o (resolve_tac [refl, conjI] ORELSE' atac));
 
 fun mk_coinduct_distinct_ctrs_tac ctxt discs discs' =
-  hyp_subst_tac ctxt THEN' REPEAT o etac conjE THEN'
-  full_simp_tac (ss_only (refl :: no_refl (union Thm.eq_thm discs discs') @ basic_simp_thms) ctxt);
+  let
+    val discs'' = map (perhaps (try (fn th => th RS @{thm notnotD}))) (discs @ discs')
+      |> distinct Thm.eq_thm_prop;
+  in
+    hyp_subst_tac ctxt THEN' REPEAT o etac conjE THEN'
+    full_simp_tac (ss_only (refl :: no_refl discs'' @ basic_simp_thms) ctxt)
+  end;
 
 fun mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn kk n pre_rel_def dtor_ctor exhaust ctr_defs
     discss selss =
--- a/src/HOL/BNF/Tools/bnf_fp_n2m.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Tools/bnf_fp_n2m.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -23,7 +23,7 @@
 open BNF_FP_N2M_Tactics
 
 fun force_typ ctxt T =
-  map_types Type_Infer.paramify_vars 
+  map_types Type_Infer.paramify_vars
   #> Type.constraint T
   #> Syntax.check_term ctxt
   #> singleton (Variable.polymorphic ctxt);
@@ -99,10 +99,6 @@
     val fp_nesty_bnfss = fp_bnfs :: nesty_bnfss;
     val fp_nesty_bnfs = distinct eq_bnf (flat fp_nesty_bnfss);
 
-    fun abstract t =
-      let val Ts = Term.add_frees t [];
-      in fold_rev Term.absfree (filter (member op = Ts) phis') t end;
-
     val rels =
       let
         fun find_rel T As Bs = fp_nesty_bnfss
@@ -121,10 +117,11 @@
               in
                 Term.list_comb (rel, rels)
               end
-          | mk_rel (T as TFree _) _ = nth phis (find_index (curry op = T) As)
+          | mk_rel (T as TFree _) _ = (nth phis (find_index (curry op = T) As)
+              handle General.Subscript => HOLogic.eq_const T)
           | mk_rel _ _ = raise Fail "fpTs contains schematic type variables";
       in
-        map2 (abstract oo mk_rel) fpTs fpTs'
+        map2 (fold_rev Term.absfree phis' oo mk_rel) fpTs fpTs'
       end;
 
     val pre_rels = map2 (fn Ds => mk_rel_of_bnf Ds (As @ fpTs) (Bs @ fpTs')) Dss bnfs;
@@ -224,7 +221,7 @@
         fun mk_s TU' =
           let
             val i = find_index (fn T => co_alg_argT TU' = T) Xs;
-            val sF = co_alg_funT TU'; 
+            val sF = co_alg_funT TU';
             val F = nth iter_preTs i;
             val s = nth iter_strs i;
           in
@@ -238,7 +235,7 @@
                   |> force_typ names_lthy smapT
                   |> hidden_to_unit;
                 val smap_argTs = strip_typeN live (fastype_of smap) |> fst;
-                fun mk_smap_arg TU =              
+                fun mk_smap_arg TU =
                   (if domain_type TU = range_type TU then
                     HOLogic.id_const (domain_type TU)
                   else if is_rec then
@@ -265,7 +262,7 @@
       in
         (case b_opt of
           NONE => ((t, Drule.dummy_thm), lthy)
-        | SOME b => Local_Theory.define ((b, NoSyn), ((Thm.def_binding b, []), 
+        | SOME b => Local_Theory.define ((b, NoSyn), ((Binding.conceal (Thm.def_binding b), []),
             fold_rev Term.absfree (if is_rec then rec_strs' else fold_strs') t)) lthy |>> apsnd snd)
       end;
 
@@ -376,6 +373,6 @@
        |> morph_fp_result (Morphism.term_morphism (singleton (Variable.polymorphic lthy))));
   in
     (fp_res, lthy)
-  end
+  end;
 
 end;
--- a/src/HOL/BNF/Tools/bnf_fp_n2m_sugar.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Tools/bnf_fp_n2m_sugar.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -7,14 +7,16 @@
 
 signature BNF_FP_N2M_SUGAR =
 sig
-  val mutualize_fp_sugars: bool -> BNF_FP_Util.fp_kind -> binding list -> typ list ->
-    (term -> int list) -> term list list list list -> BNF_FP_Def_Sugar.fp_sugar list ->
-    local_theory ->
+  val unfold_let: term -> term
+  val dest_map: Proof.context -> string -> term -> term * term list
+
+  val mutualize_fp_sugars: BNF_FP_Util.fp_kind -> binding list -> typ list -> (term -> int list) ->
+    term list list list list -> BNF_FP_Def_Sugar.fp_sugar list -> local_theory ->
     (BNF_FP_Def_Sugar.fp_sugar list
      * (BNF_FP_Def_Sugar.lfp_sugar_thms option * BNF_FP_Def_Sugar.gfp_sugar_thms option))
     * local_theory
-  val pad_and_indexify_calls: BNF_FP_Def_Sugar.fp_sugar list -> int ->
-    (term * term list list) list list -> term list list list list
+  val indexify_callsss: BNF_FP_Def_Sugar.fp_sugar -> (term * term list list) list ->
+    term list list list
   val nested_to_mutual_fps: BNF_FP_Util.fp_kind -> binding list -> typ list -> (term -> int list) ->
     (term * term list list) list list -> local_theory ->
     (typ list * int list * BNF_FP_Def_Sugar.fp_sugar list
@@ -34,171 +36,246 @@
 
 val n2mN = "n2m_"
 
-(* TODO: test with sort constraints on As *)
-(* TODO: use right sorting order for "fp_sort" w.r.t. original BNFs (?) -- treat new variables
-   as deads? *)
-fun mutualize_fp_sugars mutualize fp bs fpTs get_indices callssss fp_sugars0 no_defs_lthy0 =
-  if mutualize orelse has_duplicates (op =) fpTs then
-    let
-      val thy = Proof_Context.theory_of no_defs_lthy0;
+type n2m_sugar = fp_sugar list * (lfp_sugar_thms option * gfp_sugar_thms option);
+
+structure Data = Generic_Data
+(
+  type T = n2m_sugar Typtab.table;
+  val empty = Typtab.empty;
+  val extend = I;
+  val merge = Typtab.merge (eq_fst (eq_list eq_fp_sugar));
+);
 
-      val qsotm = quote o Syntax.string_of_term no_defs_lthy0;
+fun morph_n2m_sugar phi (fp_sugars, (lfp_sugar_thms_opt, gfp_sugar_thms_opt)) =
+  (map (morph_fp_sugar phi) fp_sugars,
+   (Option.map (morph_lfp_sugar_thms phi) lfp_sugar_thms_opt,
+    Option.map (morph_gfp_sugar_thms phi) gfp_sugar_thms_opt));
+
+val transfer_n2m_sugar =
+  morph_n2m_sugar o Morphism.thm_morphism o Thm.transfer o Proof_Context.theory_of;
 
-      fun heterogeneous_call t = error ("Heterogeneous recursive call: " ^ qsotm t);
-      fun incompatible_calls t1 t2 =
-        error ("Incompatible recursive calls: " ^ qsotm t1 ^ " vs. " ^ qsotm t2);
+fun n2m_sugar_of ctxt =
+  Typtab.lookup (Data.get (Context.Proof ctxt))
+  #> Option.map (transfer_n2m_sugar ctxt);
 
-      val b_names = map Binding.name_of bs;
-      val fp_b_names = map base_name_of_typ fpTs;
+fun register_n2m_sugar key n2m_sugar =
+  Local_Theory.declaration {syntax = false, pervasive = false}
+    (fn phi => Data.map (Typtab.default (key, morph_n2m_sugar phi n2m_sugar)));
 
-      val nn = length fpTs;
+fun unfold_let (Const (@{const_name Let}, _) $ arg1 $ arg2) = unfold_let (betapply (arg2, arg1))
+  | unfold_let (Const (@{const_name prod_case}, _) $ t) =
+    (case unfold_let t of
+      t' as Abs (s1, T1, Abs (s2, T2, _)) =>
+      let
+        val x = (s1 ^ s2, Term.maxidx_of_term t + 1);
+        val v = Var (x, HOLogic.mk_prodT (T1, T2));
+      in
+        lambda v (unfold_let (betapplys (t', [HOLogic.mk_fst v, HOLogic.mk_snd v])))
+      end
+    | _ => t)
+  | unfold_let (t $ u) = betapply (unfold_let t, unfold_let u)
+  | unfold_let (Abs (s, T, t)) = Abs (s, T, unfold_let t)
+  | unfold_let t = t;
 
-      fun target_ctr_sugar_of_fp_sugar fpT ({T, index, ctr_sugars, ...} : fp_sugar) =
-        let
-          val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (T, fpT) Vartab.empty) [];
-          val phi = Morphism.term_morphism (Term.subst_TVars rho);
-        in
-          morph_ctr_sugar phi (nth ctr_sugars index)
-        end;
-
-      val ctr_defss = map (of_fp_sugar #ctr_defss) fp_sugars0;
-      val mapss = map (of_fp_sugar #mapss) fp_sugars0;
-      val ctr_sugars0 = map2 target_ctr_sugar_of_fp_sugar fpTs fp_sugars0;
-
-      val ctrss = map #ctrs ctr_sugars0;
-      val ctr_Tss = map (map fastype_of) ctrss;
+fun mk_map_pattern ctxt s =
+  let
+    val bnf = the (bnf_of ctxt s);
+    val mapx = map_of_bnf bnf;
+    val live = live_of_bnf bnf;
+    val (f_Ts, _) = strip_typeN live (fastype_of mapx);
+    val fs = map_index (fn (i, T) => Var (("?f", i), T)) f_Ts;
+  in
+    (mapx, betapplys (mapx, fs))
+  end;
 
-      val As' = fold (fold Term.add_tfreesT) ctr_Tss [];
-      val As = map TFree As';
+fun dest_map ctxt s call =
+  let
+    val (map0, pat) = mk_map_pattern ctxt s;
+    val (_, tenv) = fo_match ctxt call pat;
+  in
+    (map0, Vartab.fold_rev (fn (_, (_, f)) => cons f) tenv [])
+  end;
+
+fun dest_abs_or_applied_map _ _ (Abs (_, _, t)) = (Term.dummy, [t])
+  | dest_abs_or_applied_map ctxt s (t1 $ _) = dest_map ctxt s t1;
 
-      val ((Cs, Xs), no_defs_lthy) =
-        no_defs_lthy0
-        |> fold Variable.declare_typ As
-        |> mk_TFrees nn
-        ||>> variant_tfrees fp_b_names;
+fun map_partition f xs =
+  fold_rev (fn x => fn (ys, (good, bad)) =>
+      case f x of SOME y => (y :: ys, (x :: good, bad)) | NONE => (ys, (good, x :: bad)))
+    xs ([], ([], []));
 
-      fun freeze_fp_default (T as Type (s, Ts)) =
-          (case find_index (curry (op =) T) fpTs of
-            ~1 => Type (s, map freeze_fp_default Ts)
-          | kk => nth Xs kk)
-        | freeze_fp_default T = T;
+fun key_of_fp_eqs fp fpTs fp_eqs =
+  Type (fp_case fp "l" "g", fpTs @ maps (fn (x, T) => [TFree x, T]) fp_eqs);
+
+(* TODO: test with sort constraints on As *)
+fun mutualize_fp_sugars fp bs fpTs get_indices callssss fp_sugars0 no_defs_lthy0 =
+  let
+    val thy = Proof_Context.theory_of no_defs_lthy0;
+
+    val qsotm = quote o Syntax.string_of_term no_defs_lthy0;
 
-      fun get_indices_checked call =
-        (case get_indices call of
-          _ :: _ :: _ => heterogeneous_call call
-        | kks => kks);
+    fun incompatible_calls t1 t2 =
+      error ("Incompatible " ^ co_prefix fp ^ "recursive calls: " ^ qsotm t1 ^ " vs. " ^ qsotm t2);
+
+    val b_names = map Binding.name_of bs;
+    val fp_b_names = map base_name_of_typ fpTs;
+
+    val nn = length fpTs;
 
-      fun freeze_fp calls (T as Type (s, Ts)) =
-          (case map_filter (try (snd o dest_map no_defs_lthy s)) calls of
-            [] =>
-            (case union (op = o pairself fst)
-                (maps (fn call => map (rpair call) (get_indices_checked call)) calls) [] of
-              [] => freeze_fp_default T
-            | [(kk, _)] => nth Xs kk
-            | (_, call1) :: (_, call2) :: _ => incompatible_calls call1 call2)
-          | callss =>
-            Type (s, map2 freeze_fp (flatten_type_args_of_bnf (the (bnf_of no_defs_lthy s)) []
-              (transpose callss)) Ts))
-        | freeze_fp _ T = T;
+    fun target_ctr_sugar_of_fp_sugar fpT ({T, index, ctr_sugars, ...} : fp_sugar) =
+      let
+        val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (T, fpT) Vartab.empty) [];
+        val phi = Morphism.term_morphism (Term.subst_TVars rho);
+      in
+        morph_ctr_sugar phi (nth ctr_sugars index)
+      end;
 
-      val ctr_Tsss = map (map binder_types) ctr_Tss;
-      val ctrXs_Tsss = map2 (map2 (map2 freeze_fp)) callssss ctr_Tsss;
-      val ctrXs_sum_prod_Ts = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctrXs_Tsss;
-      val Ts = map (body_type o hd) ctr_Tss;
+    val ctr_defss = map (of_fp_sugar #ctr_defss) fp_sugars0;
+    val mapss = map (of_fp_sugar #mapss) fp_sugars0;
+    val ctr_sugars0 = map2 target_ctr_sugar_of_fp_sugar fpTs fp_sugars0;
+
+    val ctrss = map #ctrs ctr_sugars0;
+    val ctr_Tss = map (map fastype_of) ctrss;
+
+    val As' = fold (fold Term.add_tfreesT) ctr_Tss [];
+    val As = map TFree As';
 
-      val ns = map length ctr_Tsss;
-      val kss = map (fn n => 1 upto n) ns;
-      val mss = map (map length) ctr_Tsss;
+    val ((Cs, Xs), no_defs_lthy) =
+      no_defs_lthy0
+      |> fold Variable.declare_typ As
+      |> mk_TFrees nn
+      ||>> variant_tfrees fp_b_names;
 
-      val fp_eqs = map dest_TFree Xs ~~ ctrXs_sum_prod_Ts;
+    fun check_call_dead live_call call =
+      if null (get_indices call) then () else incompatible_calls live_call call;
 
-      val base_fp_names = Name.variant_list [] fp_b_names;
-      val fp_bs = map2 (fn b_name => fn base_fp_name =>
-          Binding.qualify true b_name (Binding.name (n2mN ^ base_fp_name)))
-        b_names base_fp_names;
+    fun freeze_fpTs_simple (T as Type (s, Ts)) =
+        (case find_index (curry (op =) T) fpTs of
+          ~1 => Type (s, map freeze_fpTs_simple Ts)
+        | kk => nth Xs kk)
+      | freeze_fpTs_simple T = T;
 
-      val (pre_bnfs, (fp_res as {xtor_co_iterss = xtor_co_iterss0, xtor_co_induct,
-             dtor_injects, dtor_ctors, xtor_co_iter_thmss, ...}, lthy)) =
-        fp_bnf (construct_mutualized_fp fp fpTs fp_sugars0) fp_bs As' fp_eqs no_defs_lthy;
-
-      val nesting_bnfs = nesty_bnfs lthy ctrXs_Tsss As;
-      val nested_bnfs = nesty_bnfs lthy ctrXs_Tsss Xs;
+    fun freeze_fpTs_map (callss, (live_call :: _, dead_calls)) s Ts =
+      (List.app (check_call_dead live_call) dead_calls;
+       Type (s, map2 freeze_fpTs (flatten_type_args_of_bnf (the (bnf_of no_defs_lthy s)) []
+         (transpose callss)) Ts))
+    and freeze_fpTs calls (T as Type (s, Ts)) =
+        (case map_partition (try (snd o dest_map no_defs_lthy s)) calls of
+          ([], _) =>
+          (case map_partition (try (snd o dest_abs_or_applied_map no_defs_lthy s)) calls of
+            ([], _) => freeze_fpTs_simple T
+          | callsp => freeze_fpTs_map callsp s Ts)
+        | callsp => freeze_fpTs_map callsp s Ts)
+      | freeze_fpTs _ T = T;
 
-      val ((xtor_co_iterss, iters_args_types, coiters_args_types), _) =
-        mk_co_iters_prelims fp ctr_Tsss fpTs Cs ns mss xtor_co_iterss0 lthy;
+    val ctr_Tsss = map (map binder_types) ctr_Tss;
+    val ctrXs_Tsss = map2 (map2 (map2 freeze_fpTs)) callssss ctr_Tsss;
+    val ctrXs_sum_prod_Ts = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctrXs_Tsss;
+    val Ts = map (body_type o hd) ctr_Tss;
 
-      fun mk_binding b suf = Binding.suffix_name ("_" ^ suf) b;
+    val ns = map length ctr_Tsss;
+    val kss = map (fn n => 1 upto n) ns;
+    val mss = map (map length) ctr_Tsss;
 
-      val ((co_iterss, co_iter_defss), lthy) =
-        fold_map2 (fn b =>
-          (if fp = Least_FP then define_iters [foldN, recN] (the iters_args_types)
-           else define_coiters [unfoldN, corecN] (the coiters_args_types))
-            (mk_binding b) fpTs Cs) fp_bs xtor_co_iterss lthy
-        |>> split_list;
+    val fp_eqs = map dest_TFree Xs ~~ ctrXs_sum_prod_Ts;
+    val key = key_of_fp_eqs fp fpTs fp_eqs;
+  in
+    (case n2m_sugar_of no_defs_lthy key of
+      SOME n2m_sugar => (n2m_sugar, no_defs_lthy)
+    | NONE =>
+      let
+        val base_fp_names = Name.variant_list [] fp_b_names;
+        val fp_bs = map2 (fn b_name => fn base_fp_name =>
+            Binding.qualify true b_name (Binding.name (n2mN ^ base_fp_name)))
+          b_names base_fp_names;
 
-      val rho = tvar_subst thy Ts fpTs;
-      val ctr_sugar_phi =
-        Morphism.compose (Morphism.typ_morphism (Term.typ_subst_TVars rho))
-          (Morphism.term_morphism (Term.subst_TVars rho));
-      val inst_ctr_sugar = morph_ctr_sugar ctr_sugar_phi;
+        val (pre_bnfs, (fp_res as {xtor_co_iterss = xtor_co_iterss0, xtor_co_induct, dtor_injects,
+               dtor_ctors, xtor_co_iter_thmss, ...}, lthy)) =
+          fp_bnf (construct_mutualized_fp fp fpTs fp_sugars0) fp_bs As' fp_eqs no_defs_lthy;
+
+        val nesting_bnfs = nesty_bnfs lthy ctrXs_Tsss As;
+        val nested_bnfs = nesty_bnfs lthy ctrXs_Tsss Xs;
 
-      val ctr_sugars = map inst_ctr_sugar ctr_sugars0;
+        val ((xtor_co_iterss, iters_args_types, coiters_args_types), _) =
+          mk_co_iters_prelims fp ctr_Tsss fpTs Cs ns mss xtor_co_iterss0 lthy;
+
+        fun mk_binding b suf = Binding.suffix_name ("_" ^ suf) b;
+
+        val ((co_iterss, co_iter_defss), lthy) =
+          fold_map2 (fn b =>
+            (if fp = Least_FP then define_iters [foldN, recN] (the iters_args_types)
+             else define_coiters [unfoldN, corecN] (the coiters_args_types))
+              (mk_binding b) fpTs Cs) fp_bs xtor_co_iterss lthy
+          |>> split_list;
+
+        val rho = tvar_subst thy Ts fpTs;
+        val ctr_sugar_phi = Morphism.compose (Morphism.typ_morphism (Term.typ_subst_TVars rho))
+            (Morphism.term_morphism (Term.subst_TVars rho));
+        val inst_ctr_sugar = morph_ctr_sugar ctr_sugar_phi;
+
+        val ctr_sugars = map inst_ctr_sugar ctr_sugars0;
 
-      val ((co_inducts, 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 info => (SOME info, NONE))
-        else
-          derive_coinduct_coiters_thms_for_types pre_bnfs (the coiters_args_types) xtor_co_induct
-            dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs kss mss ns ctr_defss
-            ctr_sugars co_iterss co_iter_defss (Proof_Context.export lthy no_defs_lthy) lthy
-          |> `(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))
-          ||> (fn info => (NONE, SOME info));
+        val ((co_inducts, 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 info => (SOME info, NONE))
+          else
+            derive_coinduct_coiters_thms_for_types pre_bnfs (the coiters_args_types) xtor_co_induct
+              dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss
+              ns ctr_defss ctr_sugars co_iterss co_iter_defss
+              (Proof_Context.export lthy no_defs_lthy) lthy
+            |> `(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))
+            ||> (fn info => (NONE, SOME info));
 
-      val phi = Proof_Context.export_morphism no_defs_lthy no_defs_lthy0;
+        val phi = Proof_Context.export_morphism no_defs_lthy no_defs_lthy0;
 
-      fun mk_target_fp_sugar (kk, T) =
-        {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 = 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]}
-        |> morph_fp_sugar phi;
-    in
-      ((map_index mk_target_fp_sugar fpTs, fp_sugar_thms), lthy)
-    end
-  else
-    (* TODO: reorder hypotheses and predicates in (co)induction rules? *)
-    ((fp_sugars0, (NONE, NONE)), no_defs_lthy0);
+        fun mk_target_fp_sugar (kk, T) =
+          {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 = 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]}
+          |> morph_fp_sugar phi;
+
+        val n2m_sugar = (map_index mk_target_fp_sugar fpTs, fp_sugar_thms);
+      in
+        (n2m_sugar, lthy |> register_n2m_sugar key n2m_sugar)
+      end)
+  end;
 
 fun indexify_callsss fp_sugar callsss =
   let
     val {ctrs, ...} = of_fp_sugar #ctr_sugars fp_sugar;
-    fun do_ctr ctr =
+    fun indexify_ctr ctr =
       (case AList.lookup Term.aconv_untyped callsss ctr of
         NONE => replicate (num_binder_types (fastype_of ctr)) []
-      | SOME callss => map (map Envir.beta_eta_contract) callss);
+      | SOME callss => map (map (Envir.beta_eta_contract o unfold_let)) callss);
   in
-    map do_ctr ctrs
+    map indexify_ctr ctrs
   end;
 
-fun pad_and_indexify_calls fp_sugars0 = map2 indexify_callsss fp_sugars0 oo pad_list [];
+fun retypargs tyargs (Type (s, _)) = Type (s, tyargs);
+
+fun fold_subtype_pairs f (T as Type (s, Ts), U as Type (s', Us)) =
+    f (T, U) #> (if s = s' then fold (fold_subtype_pairs f) (Ts ~~ Us) else I)
+  | fold_subtype_pairs f TU = f TU;
 
 fun nested_to_mutual_fps fp actual_bs actual_Ts get_indices actual_callssss0 lthy =
   let
     val qsoty = quote o Syntax.string_of_typ lthy;
     val qsotys = space_implode " or " o map qsoty;
 
+    fun duplicate_datatype T = error (qsoty T ^ " is not mutually recursive with itself");
     fun not_co_datatype0 T = error (qsoty T ^ " is not a " ^ co_prefix fp ^ "datatype");
     fun not_co_datatype (T as Type (s, _)) =
         if fp = Least_FP andalso
@@ -208,32 +285,80 @@
           not_co_datatype0 T
       | not_co_datatype T = not_co_datatype0 T;
     fun not_mutually_nested_rec Ts1 Ts2 =
-      error (qsotys Ts1 ^ " is neither mutually recursive with nor nested recursive via " ^
-        qsotys Ts2);
+      error (qsotys Ts1 ^ " is neither mutually recursive with " ^ qsotys Ts2 ^
+        " nor nested recursive via " ^ qsotys Ts2);
+
+    val _ = (case Library.duplicates (op =) actual_Ts of [] => () | T :: _ => duplicate_datatype T);
 
-    val perm_actual_Ts as Type (_, ty_args0) :: _ =
-      sort (int_ord o pairself Term.size_of_typ) actual_Ts;
+    val perm_actual_Ts =
+      sort (prod_ord int_ord Term_Ord.typ_ord o pairself (`Term.size_of_typ)) actual_Ts;
+
+    fun the_ctrs_of (Type (s, Ts)) = map (mk_ctr Ts) (#ctrs (the (ctr_sugar_of lthy s)));
+
+    fun the_fp_sugar_of (T as Type (T_name, _)) =
+      (case fp_sugar_of lthy T_name of
+        SOME (fp_sugar as {fp = fp', ...}) => if fp = fp' then fp_sugar else not_co_datatype T
+      | NONE => not_co_datatype T);
 
-    fun check_enrich_with_mutuals _ [] = []
-      | check_enrich_with_mutuals seen ((T as Type (T_name, ty_args)) :: Ts) =
-        (case fp_sugar_of lthy T_name of
-          SOME ({fp = fp', fp_res = {Ts = Ts', ...}, ...}) =>
-          if fp = fp' then
+    fun gen_rhss_in gen_Ts rho subTs =
+      let
+        fun maybe_insert (T, Type (_, gen_tyargs)) =
+            if member (op =) subTs T then insert (op =) gen_tyargs else I
+          | maybe_insert _ = I;
+
+        val ctrs = maps the_ctrs_of gen_Ts;
+        val gen_ctr_Ts = maps (binder_types o fastype_of) ctrs;
+        val ctr_Ts = map (Term.typ_subst_atomic rho) gen_ctr_Ts;
+      in
+        fold (fold_subtype_pairs maybe_insert) (ctr_Ts ~~ gen_ctr_Ts) []
+      end;
+
+    fun gather_types _ _ num_groups seen gen_seen [] = (num_groups, seen, gen_seen)
+      | gather_types lthy rho num_groups seen gen_seen ((T as Type (_, tyargs)) :: Ts) =
+        let
+          val {fp_res = {Ts = mutual_Ts0, ...}, ...} = the_fp_sugar_of T;
+          val mutual_Ts = map (retypargs tyargs) mutual_Ts0;
+
+          val _ = seen = [] orelse exists (exists_subtype_in seen) mutual_Ts orelse
+            not_mutually_nested_rec mutual_Ts seen;
+
+          fun fresh_tyargs () =
             let
-              val mutual_Ts = map (fn Type (s, _) => Type (s, ty_args)) Ts';
-              val _ =
-                seen = [] orelse exists (exists_subtype_in seen) mutual_Ts orelse
-                not_mutually_nested_rec mutual_Ts seen;
-              val (seen', Ts') = List.partition (member (op =) mutual_Ts) Ts;
+              (* The name "'z" is unlikely to clash with the context, yielding more cache hits. *)
+              val (gen_tyargs, lthy') =
+                variant_tfrees (replicate (length tyargs) "z") lthy
+                |>> map Logic.varifyT_global;
+              val rho' = (gen_tyargs ~~ tyargs) @ rho;
             in
-              mutual_Ts @ check_enrich_with_mutuals (seen @ T :: seen') Ts'
-            end
-          else
-            not_co_datatype T
-        | NONE => not_co_datatype T)
-      | check_enrich_with_mutuals _ (T :: _) = not_co_datatype T;
+              (rho', gen_tyargs, gen_seen, lthy')
+            end;
 
-    val perm_Ts = check_enrich_with_mutuals [] perm_actual_Ts;
+          val (rho', gen_tyargs, gen_seen', lthy') =
+            if exists (exists_subtype_in seen) mutual_Ts then
+              (case gen_rhss_in gen_seen rho mutual_Ts of
+                [] => fresh_tyargs ()
+              | gen_tyargss as gen_tyargs :: gen_tyargss_tl =>
+                let
+                  val unify_pairs = split_list (maps (curry (op ~~) gen_tyargs) gen_tyargss_tl);
+                  val mgu = Type.raw_unifys unify_pairs Vartab.empty;
+                  val gen_tyargs' = map (Envir.subst_type mgu) gen_tyargs;
+                  val gen_seen' = map (Envir.subst_type mgu) gen_seen;
+                in
+                  (rho, gen_tyargs', gen_seen', lthy)
+                end)
+            else
+              fresh_tyargs ();
+
+          val gen_mutual_Ts = map (retypargs gen_tyargs) mutual_Ts0;
+          val Ts' = filter_out (member (op =) mutual_Ts) Ts;
+        in
+          gather_types lthy' rho' (num_groups + 1) (seen @ mutual_Ts) (gen_seen' @ gen_mutual_Ts)
+            Ts'
+        end
+      | gather_types _ _ _ _ _ (T :: _) = not_co_datatype T;
+
+    val (num_groups, perm_Ts, perm_gen_Ts) = gather_types lthy [] 0 [] [] perm_actual_Ts;
+    val perm_frozen_gen_Ts = map Logic.unvarifyT_global perm_gen_Ts;
 
     val missing_Ts = perm_Ts |> subtract (op =) actual_Ts;
     val Ts = actual_Ts @ missing_Ts;
@@ -241,6 +366,8 @@
     val nn = length Ts;
     val kks = 0 upto nn - 1;
 
+    val callssss0 = pad_list [] nn actual_callssss0;
+
     val common_name = mk_common_name (map Binding.name_of actual_bs);
     val bs = pad_list (Binding.name common_name) nn actual_bs;
 
@@ -249,16 +376,19 @@
 
     val perm_bs = permute bs;
     val perm_kks = permute kks;
+    val perm_callssss0 = permute callssss0;
     val perm_fp_sugars0 = map (the o fp_sugar_of lthy o fst o dest_Type) perm_Ts;
 
-    val mutualize = exists (fn Type (_, ty_args) => ty_args <> ty_args0) Ts;
-    val perm_callssss = pad_and_indexify_calls perm_fp_sugars0 nn actual_callssss0;
+    val perm_callssss = map2 indexify_callsss perm_fp_sugars0 perm_callssss0;
 
     val get_perm_indices = map (fn kk => find_index (curry (op =) kk) perm_kks) o get_indices;
 
     val ((perm_fp_sugars, fp_sugar_thms), lthy) =
-      mutualize_fp_sugars mutualize fp perm_bs perm_Ts get_perm_indices perm_callssss
-        perm_fp_sugars0 lthy;
+      if num_groups > 1 then
+        mutualize_fp_sugars fp perm_bs perm_frozen_gen_Ts get_perm_indices perm_callssss
+          perm_fp_sugars0 lthy
+      else
+        ((perm_fp_sugars0, (NONE, NONE)), lthy);
 
     val fp_sugars = unpermute perm_fp_sugars;
   in
--- a/src/HOL/BNF/Tools/bnf_fp_rec_sugar.ML	Mon Nov 11 17:34:44 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,986 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_fp_rec_sugar.ML
-    Author:     Lorenz Panny, TU Muenchen
-    Copyright   2013
-
-Recursor and corecursor sugar.
-*)
-
-signature BNF_FP_REC_SUGAR =
-sig
-  val add_primrec: (binding * typ option * mixfix) list ->
-    (Attrib.binding * term) list -> local_theory -> (term list * thm list list) * local_theory
-  val add_primrec_cmd: (binding * string option * mixfix) list ->
-    (Attrib.binding * string) list -> local_theory -> (term list * thm list list) * local_theory
-  val add_primrec_global: (binding * typ option * mixfix) list ->
-    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
-  val add_primrec_overloaded: (string * (string * typ) * bool) list ->
-    (binding * typ option * mixfix) list ->
-    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
-  val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
-    local_theory -> (string list * (term list * (int list list * thm list list))) * local_theory
-  val add_primcorecursive_cmd: bool ->
-    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
-    Proof.context -> Proof.state
-  val add_primcorec_cmd: bool ->
-    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
-    local_theory -> local_theory
-end;
-
-structure BNF_FP_Rec_Sugar : BNF_FP_REC_SUGAR =
-struct
-
-open BNF_Util
-open BNF_FP_Util
-open BNF_FP_Rec_Sugar_Util
-open BNF_FP_Rec_Sugar_Tactics
-
-val codeN = "code"
-val ctrN = "ctr"
-val discN = "disc"
-val selN = "sel"
-
-val nitpick_attrs = @{attributes [nitpick_simp]};
-val simp_attrs = @{attributes [simp]};
-val code_nitpick_attrs = Code.add_default_eqn_attrib :: nitpick_attrs;
-val code_nitpick_simp_attrs = Code.add_default_eqn_attrib :: nitpick_attrs @ simp_attrs;
-
-exception Primrec_Error of string * term list;
-
-fun primrec_error str = raise Primrec_Error (str, []);
-fun primrec_error_eqn str eqn = raise Primrec_Error (str, [eqn]);
-fun primrec_error_eqns str eqns = raise Primrec_Error (str, eqns);
-
-fun finds eq = fold_map (fn x => List.partition (curry eq x) #>> pair x);
-
-val free_name = try (fn Free (v, _) => v);
-val const_name = try (fn Const (v, _) => v);
-val undef_const = Const (@{const_name undefined}, dummyT);
-
-fun permute_args n t = list_comb (t, map Bound (0 :: (n downto 1)))
-  |> fold (K (Term.abs (Name.uu, dummyT))) (0 upto n);
-val abs_tuple = HOLogic.tupled_lambda o HOLogic.mk_tuple;
-fun drop_All t = subst_bounds (strip_qnt_vars @{const_name all} t |> map Free |> rev,
-  strip_qnt_body @{const_name all} t)
-fun abstract vs =
-  let fun a n (t $ u) = a n t $ a n u
-        | a n (Abs (v, T, b)) = Abs (v, T, a (n + 1) b)
-        | a n t = let val idx = find_index (equal t) vs in
-            if idx < 0 then t else Bound (n + idx) end
-  in a 0 end;
-fun mk_prod1 Ts (t, u) = HOLogic.pair_const (fastype_of1 (Ts, t)) (fastype_of1 (Ts, u)) $ t $ u;
-fun mk_tuple1 Ts = the_default HOLogic.unit o try (foldr1 (mk_prod1 Ts));
-
-fun get_indices fixes t = map (fst #>> Binding.name_of #> Free) fixes
-  |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
-  |> map_filter I;
-
-
-(* Primrec *)
-
-type eqn_data = {
-  fun_name: string,
-  rec_type: typ,
-  ctr: term,
-  ctr_args: term list,
-  left_args: term list,
-  right_args: term list,
-  res_type: typ,
-  rhs_term: term,
-  user_eqn: term
-};
-
-fun dissect_eqn lthy fun_names eqn' =
-  let
-    val eqn = drop_All eqn' |> HOLogic.dest_Trueprop
-      handle TERM _ =>
-        primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
-    val (lhs, rhs) = HOLogic.dest_eq eqn
-        handle TERM _ =>
-          primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
-    val (fun_name, args) = strip_comb lhs
-      |>> (fn x => if is_Free x then fst (dest_Free x)
-          else primrec_error_eqn "malformed function equation (does not start with free)" eqn);
-    val (left_args, rest) = take_prefix is_Free args;
-    val (nonfrees, right_args) = take_suffix is_Free rest;
-    val num_nonfrees = length nonfrees;
-    val _ = num_nonfrees = 1 orelse if num_nonfrees = 0 then
-      primrec_error_eqn "constructor pattern missing in left-hand side" eqn else
-      primrec_error_eqn "more than one non-variable argument in left-hand side" eqn;
-    val _ = member (op =) fun_names fun_name orelse
-      primrec_error_eqn "malformed function equation (does not start with function name)" eqn
-
-    val (ctr, ctr_args) = strip_comb (the_single nonfrees);
-    val _ = try (num_binder_types o fastype_of) ctr = SOME (length ctr_args) orelse
-      primrec_error_eqn "partially applied constructor in pattern" eqn;
-    val _ = let val d = duplicates (op =) (left_args @ ctr_args @ right_args) in null d orelse
-      primrec_error_eqn ("duplicate variable \"" ^ Syntax.string_of_term lthy (hd d) ^
-        "\" in left-hand side") eqn end;
-    val _ = forall is_Free ctr_args orelse
-      primrec_error_eqn "non-primitive pattern in left-hand side" eqn;
-    val _ =
-      let val b = fold_aterms (fn x as Free (v, _) =>
-        if (not (member (op =) (left_args @ ctr_args @ right_args) x) andalso
-        not (member (op =) fun_names v) andalso
-        not (Variable.is_fixed lthy v)) then cons x else I | _ => I) rhs []
-      in
-        null b orelse
-        primrec_error_eqn ("extra variable(s) in right-hand side: " ^
-          commas (map (Syntax.string_of_term lthy) b)) eqn
-      end;
-  in
-    {fun_name = fun_name,
-     rec_type = body_type (type_of ctr),
-     ctr = ctr,
-     ctr_args = ctr_args,
-     left_args = left_args,
-     right_args = right_args,
-     res_type = map fastype_of (left_args @ right_args) ---> fastype_of rhs,
-     rhs_term = rhs,
-     user_eqn = eqn'}
-  end;
-
-fun rewrite_map_arg get_ctr_pos rec_type res_type =
-  let
-    val pT = HOLogic.mk_prodT (rec_type, res_type);
-
-    val maybe_suc = Option.map (fn x => x + 1);
-    fun subst d (t as Bound d') = t |> d = SOME d' ? curry (op $) (fst_const pT)
-      | subst d (Abs (v, T, b)) = Abs (v, if d = SOME ~1 then pT else T, subst (maybe_suc d) b)
-      | subst d t =
-        let
-          val (u, vs) = strip_comb t;
-          val ctr_pos = try (get_ctr_pos o the) (free_name u) |> the_default ~1;
-        in
-          if ctr_pos >= 0 then
-            if d = SOME ~1 andalso length vs = ctr_pos then
-              list_comb (permute_args ctr_pos (snd_const pT), vs)
-            else if length vs > ctr_pos andalso is_some d
-                andalso d = try (fn Bound n => n) (nth vs ctr_pos) then
-              list_comb (snd_const pT $ nth vs ctr_pos, map (subst d) (nth_drop ctr_pos vs))
-            else
-              primrec_error_eqn ("recursive call not directly applied to constructor argument") t
-          else if d = SOME ~1 andalso const_name u = SOME @{const_name comp} then
-            list_comb (map_types (K dummyT) u, map2 subst [NONE, d] vs)
-          else
-            list_comb (u, map (subst (d |> d = SOME ~1 ? K NONE)) vs)
-        end
-  in
-    subst (SOME ~1)
-  end;
-
-fun subst_rec_calls lthy get_ctr_pos has_call ctr_args direct_calls indirect_calls t =
-  let
-    fun subst bound_Ts (Abs (v, T, b)) = Abs (v, T, subst (T :: bound_Ts) b)
-      | subst bound_Ts (t as g' $ y) =
-        let
-          val maybe_direct_y' = AList.lookup (op =) direct_calls y;
-          val maybe_indirect_y' = AList.lookup (op =) indirect_calls y;
-          val (g, g_args) = strip_comb g';
-          val ctr_pos = try (get_ctr_pos o the) (free_name g) |> the_default ~1;
-          val _ = ctr_pos < 0 orelse length g_args >= ctr_pos orelse
-            primrec_error_eqn "too few arguments in recursive call" t;
-        in
-          if not (member (op =) ctr_args y) then
-            pairself (subst bound_Ts) (g', y) |> (op $)
-          else if ctr_pos >= 0 then
-            list_comb (the maybe_direct_y', g_args)
-          else if is_some maybe_indirect_y' then
-            (if has_call g' then t else y)
-            |> massage_indirect_rec_call lthy has_call
-              (rewrite_map_arg get_ctr_pos) bound_Ts y (the maybe_indirect_y')
-            |> (if has_call g' then I else curry (op $) g')
-          else
-            t
-        end
-      | subst _ t = t
-  in
-    subst [] t
-    |> tap (fn u => has_call u andalso (* FIXME detect this case earlier *)
-      primrec_error_eqn "recursive call not directly applied to constructor argument" t)
-  end;
-
-fun build_rec_arg lthy (funs_data : eqn_data list list) has_call (ctr_spec : rec_ctr_spec)
-    (maybe_eqn_data : eqn_data option) =
-  if is_none maybe_eqn_data then undef_const else
-    let
-      val eqn_data = the maybe_eqn_data;
-      val t = #rhs_term eqn_data;
-      val ctr_args = #ctr_args eqn_data;
-
-      val calls = #calls ctr_spec;
-      val n_args = fold (curry (op +) o (fn Direct_Rec _ => 2 | _ => 1)) calls 0;
-
-      val no_calls' = tag_list 0 calls
-        |> map_filter (try (apsnd (fn No_Rec n => n | Direct_Rec (n, _) => n)));
-      val direct_calls' = tag_list 0 calls
-        |> map_filter (try (apsnd (fn Direct_Rec (_, n) => n)));
-      val indirect_calls' = tag_list 0 calls
-        |> map_filter (try (apsnd (fn Indirect_Rec n => n)));
-
-      fun make_direct_type _ = dummyT; (* FIXME? *)
-
-      val rec_res_type_list = map (fn (x :: _) => (#rec_type x, #res_type x)) funs_data;
-
-      fun make_indirect_type (Type (Tname, Ts)) = Type (Tname, Ts |> map (fn T =>
-        let val maybe_res_type = AList.lookup (op =) rec_res_type_list T in
-          if is_some maybe_res_type
-          then HOLogic.mk_prodT (T, the maybe_res_type)
-          else make_indirect_type T end))
-        | make_indirect_type T = T;
-
-      val args = replicate n_args ("", dummyT)
-        |> Term.rename_wrt_term t
-        |> map Free
-        |> fold (fn (ctr_arg_idx, arg_idx) =>
-            nth_map arg_idx (K (nth ctr_args ctr_arg_idx)))
-          no_calls'
-        |> fold (fn (ctr_arg_idx, arg_idx) =>
-            nth_map arg_idx (K (nth ctr_args ctr_arg_idx |> map_types make_direct_type)))
-          direct_calls'
-        |> fold (fn (ctr_arg_idx, arg_idx) =>
-            nth_map arg_idx (K (nth ctr_args ctr_arg_idx |> map_types make_indirect_type)))
-          indirect_calls';
-
-      val fun_name_ctr_pos_list =
-        map (fn (x :: _) => (#fun_name x, length (#left_args x))) funs_data;
-      val get_ctr_pos = try (the o AList.lookup (op =) fun_name_ctr_pos_list) #> the_default ~1;
-      val direct_calls = map (apfst (nth ctr_args) o apsnd (nth args)) direct_calls';
-      val indirect_calls = map (apfst (nth ctr_args) o apsnd (nth args)) indirect_calls';
-
-      val abstractions = args @ #left_args eqn_data @ #right_args eqn_data;
-    in
-      t
-      |> subst_rec_calls lthy get_ctr_pos has_call ctr_args direct_calls indirect_calls
-      |> fold_rev lambda abstractions
-    end;
-
-fun build_defs lthy bs mxs (funs_data : eqn_data list list) (rec_specs : rec_spec list) has_call =
-  let
-    val n_funs = length funs_data;
-
-    val ctr_spec_eqn_data_list' =
-      (take n_funs rec_specs |> map #ctr_specs) ~~ funs_data
-      |> maps (uncurry (finds (fn (x, y) => #ctr x = #ctr y))
-          ##> (fn x => null x orelse
-            primrec_error_eqns "excess equations in definition" (map #rhs_term x)) #> fst);
-    val _ = ctr_spec_eqn_data_list' |> map (fn (_, x) => length x <= 1 orelse
-      primrec_error_eqns ("multiple equations for constructor") (map #user_eqn x));
-
-    val ctr_spec_eqn_data_list =
-      ctr_spec_eqn_data_list' @ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
-
-    val recs = take n_funs rec_specs |> map #recx;
-    val rec_args = ctr_spec_eqn_data_list
-      |> sort ((op <) o pairself (#offset o fst) |> make_ord)
-      |> map (uncurry (build_rec_arg lthy funs_data has_call) o apsnd (try the_single));
-    val ctr_poss = map (fn x =>
-      if length (distinct ((op =) o pairself (length o #left_args)) x) <> 1 then
-        primrec_error ("inconstant constructor pattern position for function " ^
-          quote (#fun_name (hd x)))
-      else
-        hd x |> #left_args |> length) funs_data;
-  in
-    (recs, ctr_poss)
-    |-> map2 (fn recx => fn ctr_pos => list_comb (recx, rec_args) |> permute_args ctr_pos)
-    |> Syntax.check_terms lthy
-    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.map_name Thm.def_name b, []), t))) bs mxs
-  end;
-
-fun find_rec_calls has_call (eqn_data : eqn_data) =
-  let
-    fun find (Abs (_, _, b)) ctr_arg = find b ctr_arg
-      | find (t as _ $ _) ctr_arg =
-        let
-          val (f', args') = strip_comb t;
-          val n = find_index (equal ctr_arg) args';
-        in
-          if n < 0 then
-            find f' ctr_arg @ maps (fn x => find x ctr_arg) args'
-          else
-            let val (f, args) = chop n args' |>> curry list_comb f' in
-              if has_call f then
-                f :: maps (fn x => find x ctr_arg) args
-              else
-                find f ctr_arg @ maps (fn x => find x ctr_arg) args
-            end
-        end
-      | find _ _ = [];
-  in
-    map (find (#rhs_term eqn_data)) (#ctr_args eqn_data)
-    |> (fn [] => NONE | callss => SOME (#ctr eqn_data, callss))
-  end;
-
-fun prepare_primrec fixes specs lthy =
-  let
-    val (bs, mxs) = map_split (apfst fst) fixes;
-    val fun_names = map Binding.name_of bs;
-    val eqns_data = map (dissect_eqn lthy fun_names) specs;
-    val funs_data = eqns_data
-      |> partition_eq ((op =) o pairself #fun_name)
-      |> finds (fn (x, y) => x = #fun_name (hd y)) fun_names |> fst
-      |> map (fn (x, y) => the_single y handle List.Empty =>
-          primrec_error ("missing equations for function " ^ quote x));
-
-    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
-    val arg_Ts = map (#rec_type o hd) funs_data;
-    val res_Ts = map (#res_type o hd) funs_data;
-    val callssss = funs_data
-      |> map (partition_eq ((op =) o pairself #ctr))
-      |> map (maps (map_filter (find_rec_calls has_call)));
-
-    val ((n2m, rec_specs, _, induct_thm, induct_thms), lthy') =
-      rec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
-
-    val actual_nn = length funs_data;
-
-    val _ = let val ctrs = (maps (map #ctr o #ctr_specs) rec_specs) in
-      map (fn {ctr, user_eqn, ...} => member (op =) ctrs ctr orelse
-        primrec_error_eqn ("argument " ^ quote (Syntax.string_of_term lthy' ctr) ^
-          " is not a constructor in left-hand side") user_eqn) eqns_data end;
-
-    val defs = build_defs lthy' bs mxs funs_data rec_specs has_call;
-
-    fun prove lthy def_thms' ({ctr_specs, nested_map_idents, nested_map_comps, ...} : rec_spec)
-        (fun_data : eqn_data list) =
-      let
-        val def_thms = map (snd o snd) def_thms';
-        val simp_thmss = finds (fn (x, y) => #ctr x = #ctr y) fun_data ctr_specs
-          |> fst
-          |> map_filter (try (fn (x, [y]) =>
-            (#user_eqn x, length (#left_args x) + length (#right_args x), #rec_thm y)))
-          |> map (fn (user_eqn, num_extra_args, rec_thm) =>
-            mk_primrec_tac lthy num_extra_args nested_map_idents nested_map_comps def_thms rec_thm
-            |> K |> Goal.prove lthy [] [] user_eqn);
-        val poss = find_indices (fn (x, y) => #ctr x = #ctr y) fun_data eqns_data;
-      in
-        (poss, simp_thmss)
-      end;
-
-    val notes =
-      (if n2m then map2 (fn name => fn thm =>
-        (name, inductN, [thm], [])) fun_names (take actual_nn induct_thms) else [])
-      |> map (fn (prefix, thmN, thms, attrs) =>
-        ((Binding.qualify true prefix (Binding.name thmN), attrs), [(thms, [])]));
-
-    val common_name = mk_common_name fun_names;
-
-    val common_notes =
-      (if n2m then [(inductN, [induct_thm], [])] else [])
-      |> map (fn (thmN, thms, attrs) =>
-        ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
-  in
-    (((fun_names, defs),
-      fn lthy => fn defs =>
-        split_list (map2 (prove lthy defs) (take actual_nn rec_specs) funs_data)),
-      lthy' |> Local_Theory.notes (notes @ common_notes) |> snd)
-  end;
-
-(* primrec definition *)
-
-fun add_primrec_simple fixes ts lthy =
-  let
-    val (((names, defs), prove), lthy) = prepare_primrec fixes ts lthy
-      handle ERROR str => primrec_error str;
-  in
-    lthy
-    |> fold_map Local_Theory.define defs
-    |-> (fn defs => `(fn lthy => (names, (map fst defs, prove lthy defs))))
-  end
-  handle Primrec_Error (str, eqns) =>
-    if null eqns
-    then error ("primrec_new error:\n  " ^ str)
-    else error ("primrec_new error:\n  " ^ str ^ "\nin\n  " ^
-      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
-
-local
-
-fun gen_primrec prep_spec (raw_fixes : (binding * 'a option * mixfix) list) raw_spec lthy =
-  let
-    val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes)
-    val _ = null d orelse primrec_error ("duplicate function name(s): " ^ commas d);
-
-    val (fixes, specs) = fst (prep_spec raw_fixes raw_spec lthy);
-
-    val mk_notes =
-      flat ooo map3 (fn poss => fn prefix => fn thms =>
-        let
-          val (bs, attrss) = map_split (fst o nth specs) poss;
-          val notes =
-            map3 (fn b => fn attrs => fn thm =>
-              ((Binding.qualify false prefix b, code_nitpick_simp_attrs @ attrs), [([thm], [])]))
-            bs attrss thms;
-        in
-          ((Binding.qualify true prefix (Binding.name simpsN), []), [(thms, [])]) :: notes
-        end);
-  in
-    lthy
-    |> add_primrec_simple fixes (map snd specs)
-    |-> (fn (names, (ts, (posss, simpss))) =>
-      Spec_Rules.add Spec_Rules.Equational (ts, flat simpss)
-      #> Local_Theory.notes (mk_notes posss names simpss)
-      #>> pair ts o map snd)
-  end;
-
-in
-
-val add_primrec = gen_primrec Specification.check_spec;
-val add_primrec_cmd = gen_primrec Specification.read_spec;
-
-end;
-
-fun add_primrec_global fixes specs thy =
-  let
-    val lthy = Named_Target.theory_init thy;
-    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
-    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
-  in ((ts, simps'), Local_Theory.exit_global lthy') end;
-
-fun add_primrec_overloaded ops fixes specs thy =
-  let
-    val lthy = Overloading.overloading ops thy;
-    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
-    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
-  in ((ts, simps'), Local_Theory.exit_global lthy') end;
-
-
-
-(* Primcorec *)
-
-type co_eqn_data_disc = {
-  fun_name: string,
-  fun_T: typ,
-  fun_args: term list,
-  ctr: term,
-  ctr_no: int, (*###*)
-  disc: term,
-  prems: term list,
-  auto_gen: bool,
-  user_eqn: term
-};
-
-type co_eqn_data_sel = {
-  fun_name: string,
-  fun_T: typ,
-  fun_args: term list,
-  ctr: term,
-  sel: term,
-  rhs_term: term,
-  user_eqn: term
-};
-
-datatype co_eqn_data =
-  Disc of co_eqn_data_disc |
-  Sel of co_eqn_data_sel;
-
-fun co_dissect_eqn_disc sequential fun_names (corec_specs : corec_spec list) prems' concl
-    matchedsss =
-  let
-    fun find_subterm p = let (* FIXME \<exists>? *)
-      fun f (t as u $ v) = if p t then SOME t else merge_options (f u, f v)
-        | f t = if p t then SOME t else NONE
-      in f end;
-
-    val applied_fun = concl
-      |> find_subterm (member ((op =) o apsnd SOME) fun_names o try (fst o dest_Free o head_of))
-      |> the
-      handle Option.Option => primrec_error_eqn "malformed discriminator equation" concl;
-    val ((fun_name, fun_T), fun_args) = strip_comb applied_fun |>> dest_Free;
-    val {ctr_specs, ...} = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name);
-
-    val discs = map #disc ctr_specs;
-    val ctrs = map #ctr ctr_specs;
-    val not_disc = head_of concl = @{term Not};
-    val _ = not_disc andalso length ctrs <> 2 andalso
-      primrec_error_eqn "\<not>ed discriminator for a type with \<noteq> 2 constructors" concl;
-    val disc = find_subterm (member (op =) discs o head_of) concl;
-    val eq_ctr0 = concl |> perhaps (try (HOLogic.dest_not)) |> try (HOLogic.dest_eq #> snd)
-        |> (fn SOME t => let val n = find_index (equal t) ctrs in
-          if n >= 0 then SOME n else NONE end | _ => NONE);
-    val _ = is_some disc orelse is_some eq_ctr0 orelse
-      primrec_error_eqn "no discriminator in equation" concl;
-    val ctr_no' =
-      if is_none disc then the eq_ctr0 else find_index (equal (head_of (the disc))) discs;
-    val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
-    val ctr = #ctr (nth ctr_specs ctr_no);
-
-    val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_;
-    val matchedss = AList.lookup (op =) matchedsss fun_name |> the_default [];
-    val prems = map (abstract (List.rev fun_args)) prems';
-    val real_prems =
-      (if catch_all orelse sequential then maps negate_disj matchedss else []) @
-      (if catch_all then [] else prems);
-
-    val matchedsss' = AList.delete (op =) fun_name matchedsss
-      |> cons (fun_name, if sequential then matchedss @ [prems] else matchedss @ [real_prems]);
-
-    val user_eqn =
-      (real_prems, betapply (#disc (nth ctr_specs ctr_no), applied_fun))
-      |>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop
-      |> Logic.list_implies;
-  in
-    (Disc {
-      fun_name = fun_name,
-      fun_T = fun_T,
-      fun_args = fun_args,
-      ctr = ctr,
-      ctr_no = ctr_no,
-      disc = #disc (nth ctr_specs ctr_no),
-      prems = real_prems,
-      auto_gen = catch_all,
-      user_eqn = user_eqn
-    }, matchedsss')
-  end;
-
-fun co_dissect_eqn_sel fun_names (corec_specs : corec_spec list) eqn' of_spec eqn =
-  let
-    val (lhs, rhs) = HOLogic.dest_eq eqn
-      handle TERM _ =>
-        primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
-    val sel = head_of lhs;
-    val ((fun_name, fun_T), fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst dest_Free
-      handle TERM _ =>
-        primrec_error_eqn "malformed selector argument in left-hand side" eqn;
-    val corec_spec = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name)
-      handle Option.Option => primrec_error_eqn "malformed selector argument in left-hand side" eqn;
-    val ctr_spec =
-      if is_some of_spec
-      then the (find_first (equal (the of_spec) o #ctr) (#ctr_specs corec_spec))
-      else #ctr_specs corec_spec |> filter (exists (equal sel) o #sels) |> the_single
-        handle List.Empty => primrec_error_eqn "ambiguous selector - use \"of\"" eqn;
-    val user_eqn = drop_All eqn';
-  in
-    Sel {
-      fun_name = fun_name,
-      fun_T = fun_T,
-      fun_args = fun_args,
-      ctr = #ctr ctr_spec,
-      sel = sel,
-      rhs_term = rhs,
-      user_eqn = user_eqn
-    }
-  end;
-
-fun co_dissect_eqn_ctr sequential fun_names (corec_specs : corec_spec list) eqn' imp_prems imp_rhs
-    matchedsss =
-  let
-    val (lhs, rhs) = HOLogic.dest_eq imp_rhs;
-    val fun_name = head_of lhs |> fst o dest_Free;
-    val {ctr_specs, ...} = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name);
-    val (ctr, ctr_args) = strip_comb rhs;
-    val {disc, sels, ...} = the (find_first (equal ctr o #ctr) ctr_specs)
-      handle Option.Option => primrec_error_eqn "not a constructor" ctr;
-
-    val disc_imp_rhs = betapply (disc, lhs);
-    val (maybe_eqn_data_disc, matchedsss') = if length ctr_specs = 1
-      then (NONE, matchedsss)
-      else apfst SOME (co_dissect_eqn_disc
-          sequential fun_names corec_specs imp_prems disc_imp_rhs matchedsss);
-
-    val sel_imp_rhss = (sels ~~ ctr_args)
-      |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
-
-(*
-val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} imp_rhs ^ "\nto\n    \<cdot> " ^
- (is_some maybe_eqn_data_disc ? K (Syntax.string_of_term @{context} disc_imp_rhs ^ "\n    \<cdot> ")) "" ^
- space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_imp_rhss));
-*)
-
-    val eqns_data_sel =
-      map (co_dissect_eqn_sel fun_names corec_specs eqn' (SOME ctr)) sel_imp_rhss;
-  in
-    (the_list maybe_eqn_data_disc @ eqns_data_sel, matchedsss')
-  end;
-
-fun co_dissect_eqn sequential fun_names (corec_specs : corec_spec list) eqn' of_spec matchedsss =
-  let
-    val eqn = drop_All eqn'
-      handle TERM _ => primrec_error_eqn "malformed function equation" eqn';
-    val (imp_prems, imp_rhs) = Logic.strip_horn eqn
-      |> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop;
-
-    val head = imp_rhs
-      |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
-      |> head_of;
-
-    val maybe_rhs = imp_rhs |> perhaps (try (HOLogic.dest_not)) |> try (snd o HOLogic.dest_eq);
-
-    val discs = maps #ctr_specs corec_specs |> map #disc;
-    val sels = maps #ctr_specs corec_specs |> maps #sels;
-    val ctrs = maps #ctr_specs corec_specs |> map #ctr;
-  in
-    if member (op =) discs head orelse
-      is_some maybe_rhs andalso
-        member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then
-      co_dissect_eqn_disc sequential fun_names corec_specs imp_prems imp_rhs matchedsss
-      |>> single
-    else if member (op =) sels head then
-      ([co_dissect_eqn_sel fun_names corec_specs eqn' of_spec imp_rhs], matchedsss)
-    else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) then
-      co_dissect_eqn_ctr sequential fun_names corec_specs eqn' imp_prems imp_rhs matchedsss
-    else
-      primrec_error_eqn "malformed function equation" eqn
-  end;
-
-fun build_corec_arg_disc (ctr_specs : corec_ctr_spec list)
-    ({fun_args, ctr_no, prems, ...} : co_eqn_data_disc) =
-  if is_none (#pred (nth ctr_specs ctr_no)) then I else
-    mk_conjs prems
-    |> curry subst_bounds (List.rev fun_args)
-    |> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args)
-    |> K |> nth_map (the (#pred (nth ctr_specs ctr_no)));
-
-fun build_corec_arg_no_call (sel_eqns : co_eqn_data_sel list) sel =
-  find_first (equal sel o #sel) sel_eqns
-  |> try (fn SOME {fun_args, rhs_term, ...} => abs_tuple fun_args rhs_term)
-  |> the_default undef_const
-  |> K;
-
-fun build_corec_args_direct_call lthy has_call (sel_eqns : co_eqn_data_sel list) sel =
-  let
-    val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
-  in
-    if is_none maybe_sel_eqn then (I, I, I) else
-    let
-      val {fun_args, rhs_term, ... } = the maybe_sel_eqn;
-      fun rewrite_q _ t = if has_call t then @{term False} else @{term True};
-      fun rewrite_g _ t = if has_call t then undef_const else t;
-      fun rewrite_h bound_Ts t =
-        if has_call t then mk_tuple1 bound_Ts (snd (strip_comb t)) else undef_const;
-      fun massage f t = massage_direct_corec_call lthy has_call f [] rhs_term |> abs_tuple fun_args;
-    in
-      (massage rewrite_q,
-       massage rewrite_g,
-       massage rewrite_h)
-    end
-  end;
-
-fun build_corec_arg_indirect_call lthy has_call (sel_eqns : co_eqn_data_sel list) sel =
-  let
-    val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
-  in
-    if is_none maybe_sel_eqn then I else
-    let
-      val {fun_args, rhs_term, ...} = the maybe_sel_eqn;
-      fun rewrite bound_Ts U T (Abs (v, V, b)) = Abs (v, V, rewrite (V :: bound_Ts) U T b)
-        | rewrite bound_Ts U T (t as _ $ _) =
-          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
-              map (rewrite bound_Ts U T) vs |> chop 1 |>> HOLogic.mk_split o the_single |> list_comb
-            else
-              list_comb (rewrite bound_Ts U T u, map (rewrite bound_Ts U T) vs)
-          end
-        | rewrite _ U T t =
-          if is_Free t andalso has_call t then Inr_const U T $ HOLogic.unit else t;
-      fun massage t =
-        massage_indirect_corec_call lthy has_call rewrite [] (range_type (fastype_of t)) rhs_term
-        |> abs_tuple fun_args;
-    in
-      massage
-    end
-  end;
-
-fun build_corec_args_sel lthy has_call (all_sel_eqns : co_eqn_data_sel list)
-    (ctr_spec : corec_ctr_spec) =
-  let val sel_eqns = filter (equal (#ctr ctr_spec) o #ctr) all_sel_eqns in
-    if null sel_eqns then I else
-      let
-        val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec;
-
-        val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list;
-        val direct_calls' = map_filter (try (apsnd (fn Direct_Corec n => n))) sel_call_list;
-        val indirect_calls' = map_filter (try (apsnd (fn Indirect_Corec n => n))) sel_call_list;
-      in
-        I
-        #> fold (fn (sel, n) => nth_map n (build_corec_arg_no_call sel_eqns sel)) no_calls'
-        #> fold (fn (sel, (q, g, h)) =>
-          let val (fq, fg, fh) = build_corec_args_direct_call lthy has_call sel_eqns sel in
-            nth_map q fq o nth_map g fg o nth_map h fh end) direct_calls'
-        #> fold (fn (sel, n) => nth_map n
-          (build_corec_arg_indirect_call lthy has_call sel_eqns sel)) indirect_calls'
-      end
-  end;
-
-fun co_build_defs lthy bs mxs has_call arg_Tss (corec_specs : corec_spec list)
-    (disc_eqnss : co_eqn_data_disc list list) (sel_eqnss : co_eqn_data_sel list list) =
-  let
-    val corec_specs' = take (length bs) corec_specs;
-    val corecs = map #corec corec_specs';
-    val ctr_specss = map #ctr_specs corec_specs';
-    val corec_args = hd corecs
-      |> fst o split_last o binder_types o fastype_of
-      |> map (Const o pair @{const_name undefined})
-      |> fold2 (fold o build_corec_arg_disc) ctr_specss disc_eqnss
-      |> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss;
-    fun currys [] t = t
-      | currys Ts t = t $ mk_tuple1 (List.rev Ts) (map Bound (length Ts - 1 downto 0))
-          |> fold_rev (Term.abs o pair Name.uu) Ts;
-
-(*
-val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
- space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
-*)
-
-    val exclss' =
-      disc_eqnss
-      |> map (map (fn x => (#fun_args x, #ctr_no x, #prems x, #auto_gen x))
-        #> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs [])
-        #> maps (uncurry (map o pair)
-          #> map (fn ((fun_args, c, x, a), (_, c', y, a')) =>
-              ((c, c', a orelse a'), (x, s_not (mk_conjs y)))
-            ||> apfst (map HOLogic.mk_Trueprop) o apsnd HOLogic.mk_Trueprop
-            ||> Logic.list_implies
-            ||> curry Logic.list_all (map dest_Free fun_args))))
-  in
-    map (list_comb o rpair corec_args) corecs
-    |> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss
-    |> map2 currys arg_Tss
-    |> Syntax.check_terms lthy
-    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.map_name Thm.def_name b, []), t))) bs mxs
-    |> rpair exclss'
-  end;
-
-fun mk_real_disc_eqns fun_binding arg_Ts ({ctr_specs, ...} : corec_spec)
-    (sel_eqns : co_eqn_data_sel list) (disc_eqns : co_eqn_data_disc list) =
-  if length disc_eqns <> length ctr_specs - 1 then disc_eqns else
-    let
-      val n = 0 upto length ctr_specs
-        |> the o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns));
-      val fun_args = (try (#fun_args o hd) disc_eqns, try (#fun_args o hd) sel_eqns)
-        |> the_default (map (curry Free Name.uu) arg_Ts) o merge_options;
-      val extra_disc_eqn = {
-        fun_name = Binding.name_of fun_binding,
-        fun_T = arg_Ts ---> body_type (fastype_of (#ctr (hd ctr_specs))),
-        fun_args = fun_args,
-        ctr = #ctr (nth ctr_specs n),
-        ctr_no = n,
-        disc = #disc (nth ctr_specs n),
-        prems = maps (negate_conj o #prems) disc_eqns,
-        auto_gen = true,
-        user_eqn = undef_const};
-    in
-      chop n disc_eqns ||> cons extra_disc_eqn |> (op @)
-    end;
-
-fun add_primcorec simple sequential fixes specs of_specs lthy =
-  let
-    val (bs, mxs) = map_split (apfst fst) fixes;
-    val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list;
-
-    val callssss = []; (* FIXME *)
-
-    val ((n2m, corec_specs', _, coinduct_thm, strong_coinduct_thm, coinduct_thms,
-          strong_coinduct_thms), lthy') =
-      corec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
-
-    val actual_nn = length bs;
-    val fun_names = map Binding.name_of bs;
-    val corec_specs = take actual_nn corec_specs'; (*###*)
-
-    val eqns_data =
-      fold_map2 (co_dissect_eqn sequential fun_names corec_specs) (map snd specs) of_specs []
-      |> flat o fst;
-
-    val disc_eqnss' = map_filter (try (fn Disc x => x)) eqns_data
-      |> partition_eq ((op =) o pairself #fun_name)
-      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
-      |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
-    val _ = disc_eqnss' |> map (fn x =>
-      let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse
-        primrec_error_eqns "excess discriminator equations in definition"
-          (maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
-
-    val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data
-      |> partition_eq ((op =) o pairself #fun_name)
-      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
-      |> map (flat o snd);
-
-    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
-    val arg_Tss = map (binder_types o snd o fst) fixes;
-    val disc_eqnss = map5 mk_real_disc_eqns bs arg_Tss corec_specs sel_eqnss disc_eqnss';
-    val (defs, exclss') =
-      co_build_defs lthy' bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss;
-
-    fun excl_tac (c, c', a) =
-      if a orelse c = c' orelse sequential then
-        SOME (K (HEADGOAL (mk_primcorec_assumption_tac lthy [])))
-      else if simple then
-        SOME (K (auto_tac lthy))
-      else
-        NONE;
-
-(*
-val _ = tracing ("exclusiveness properties:\n    \<cdot> " ^
- space_implode "\n    \<cdot> " (maps (map (Syntax.string_of_term lthy o snd)) exclss'));
-*)
-
-    val exclss'' = exclss' |> map (map (fn (idx, t) =>
-      (idx, (Option.map (Goal.prove lthy [] [] t) (excl_tac idx), t))));
-    val taut_thmss = map (map (apsnd (the o fst)) o filter (is_some o fst o snd)) exclss'';
-    val (obligation_idxss, obligationss) = exclss''
-      |> map (map (apsnd (rpair [] o snd)) o filter (is_none o fst o snd))
-      |> split_list o map split_list;
-
-    fun prove thmss' def_thms' lthy =
-      let
-        val def_thms = map (snd o snd) def_thms';
-
-        val exclss' = map (op ~~) (obligation_idxss ~~ thmss');
-        fun mk_exclsss excls n =
-          (excls, map (fn k => replicate k [TrueI] @ replicate (n - k) []) (0 upto n - 1))
-          |-> fold (fn ((c, c', _), thm) => nth_map c (nth_map c' (K [thm])));
-        val exclssss = (exclss' ~~ taut_thmss |> map (op @), fun_names ~~ corec_specs)
-          |-> map2 (fn excls => fn (_, {ctr_specs, ...}) => mk_exclsss excls (length ctr_specs));
-
-        fun prove_disc ({ctr_specs, ...} : corec_spec) exclsss
-            ({fun_name, fun_T, fun_args, ctr_no, prems, ...} : co_eqn_data_disc) =
-          if Term.aconv_untyped (#disc (nth ctr_specs ctr_no), @{term "\<lambda>x. x = x"}) then [] else
-            let
-              val {disc_corec, ...} = nth ctr_specs ctr_no;
-              val k = 1 + ctr_no;
-              val m = length prems;
-              val t =
-                list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
-                |> curry betapply (#disc (nth ctr_specs ctr_no)) (*###*)
-                |> HOLogic.mk_Trueprop
-                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
-                |> curry Logic.list_all (map dest_Free fun_args);
-            in
-              mk_primcorec_disc_tac lthy def_thms disc_corec k m exclsss
-              |> K |> Goal.prove lthy [] [] t
-              |> pair (#disc (nth ctr_specs ctr_no))
-              |> single
-            end;
-
-        fun prove_sel ({nested_maps, nested_map_idents, nested_map_comps, ctr_specs, ...}
-            : corec_spec) (disc_eqns : co_eqn_data_disc list) exclsss
-            ({fun_name, fun_T, fun_args, ctr, sel, rhs_term, ...} : co_eqn_data_sel) =
-          let
-            val SOME ctr_spec = find_first (equal ctr o #ctr) ctr_specs;
-            val ctr_no = find_index (equal ctr o #ctr) ctr_specs;
-            val prems = the_default (maps (negate_conj o #prems) disc_eqns)
-                (find_first (equal ctr_no o #ctr_no) disc_eqns |> Option.map #prems);
-            val sel_corec = find_index (equal sel) (#sels ctr_spec)
-              |> nth (#sel_corecs ctr_spec);
-            val k = 1 + ctr_no;
-            val m = length prems;
-            val t =
-              list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
-              |> curry betapply sel
-              |> rpair (abstract (List.rev fun_args) rhs_term)
-              |> 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;
-          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 exclsss
-            |> K |> Goal.prove lthy [] [] t
-            |> pair sel
-          end;
-
-        fun prove_ctr disc_alist sel_alist (disc_eqns : co_eqn_data_disc list)
-            (sel_eqns : co_eqn_data_sel list) ({ctr, disc, sels, collapse, ...} : corec_ctr_spec) =
-          if not (exists (equal ctr o #ctr) disc_eqns)
-              andalso not (exists (equal ctr o #ctr) sel_eqns)
-            orelse (* don't try to prove theorems when some sel_eqns are missing *)
-              filter (equal ctr o #ctr) sel_eqns
-              |> fst o finds ((op =) o apsnd #sel) sels
-              |> exists (null o snd)
-          then [] else
-            let
-              val (fun_name, fun_T, fun_args, prems) =
-                (find_first (equal ctr o #ctr) disc_eqns, find_first (equal ctr o #ctr) sel_eqns)
-                |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #prems x))
-                ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, []))
-                |> the o merge_options;
-              val m = length prems;
-              val t = filter (equal ctr o #ctr) sel_eqns
-                |> fst o finds ((op =) o apsnd #sel) sels
-                |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x)) #-> abstract)
-                |> curry list_comb ctr
-                |> curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
-                  map Bound (length fun_args - 1 downto 0)))
-                |> HOLogic.mk_Trueprop
-                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
-                |> curry Logic.list_all (map dest_Free fun_args);
-              val maybe_disc_thm = AList.lookup (op =) disc_alist disc;
-              val sel_thms = map snd (filter (member (op =) sels o fst) sel_alist);
-            in
-              mk_primcorec_ctr_of_dtr_tac lthy m collapse maybe_disc_thm sel_thms
-              |> K |> Goal.prove lthy [] [] t
-              |> single
-            end;
-
-        val disc_alists = map3 (maps oo prove_disc) corec_specs exclssss disc_eqnss;
-        val sel_alists = map4 (map ooo prove_sel) corec_specs disc_eqnss exclssss sel_eqnss;
-
-        val disc_thmss = map (map snd) disc_alists;
-        val sel_thmss = map (map snd) sel_alists;
-        val ctr_thmss = map5 (maps oooo prove_ctr) disc_alists sel_alists disc_eqnss sel_eqnss
-          (map #ctr_specs corec_specs);
-
-        val simp_thmss = map2 append disc_thmss sel_thmss
-
-        val common_name = mk_common_name fun_names;
-
-        val notes =
-          [(coinductN, map (if n2m then single else K []) coinduct_thms, []),
-           (codeN, ctr_thmss(*FIXME*), code_nitpick_attrs),
-           (ctrN, ctr_thmss, []),
-           (discN, disc_thmss, simp_attrs),
-           (selN, sel_thmss, simp_attrs),
-           (simpsN, simp_thmss, []),
-           (strong_coinductN, map (if n2m then single else K []) strong_coinduct_thms, [])]
-          |> maps (fn (thmN, thmss, attrs) =>
-            map2 (fn fun_name => fn thms =>
-                ((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])]))
-              fun_names (take actual_nn thmss))
-          |> filter_out (null o fst o hd o snd);
-
-        val common_notes =
-          [(coinductN, if n2m then [coinduct_thm] else [], []),
-           (strong_coinductN, if n2m then [strong_coinduct_thm] else [], [])]
-          |> filter_out (null o #2)
-          |> map (fn (thmN, thms, attrs) =>
-            ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
-      in
-        lthy |> Local_Theory.notes (notes @ common_notes) |> snd
-      end;
-
-    fun after_qed thmss' = fold_map Local_Theory.define defs #-> prove thmss';
-
-    val _ = if not simple orelse forall null obligationss then () else
-      primrec_error "need exclusiveness proofs - use primcorecursive instead of primcorec";
-  in
-    if simple then
-      lthy'
-      |> after_qed (map (fn [] => []) obligationss)
-      |> pair NONE o SOME
-    else
-      lthy'
-      |> Proof.theorem NONE after_qed obligationss
-      |> Proof.refine (Method.primitive_text I)
-      |> Seq.hd
-      |> rpair NONE o SOME
-  end;
-
-fun add_primcorec_ursive_cmd simple seq (raw_fixes, raw_specs') lthy =
-  let
-    val (raw_specs, of_specs) = split_list raw_specs' ||> map (Option.map (Syntax.read_term lthy));
-    val ((fixes, specs), _) = Specification.read_spec raw_fixes raw_specs lthy;
-  in
-    add_primcorec simple seq fixes specs of_specs lthy
-    handle ERROR str => primrec_error str
-  end
-  handle Primrec_Error (str, eqns) =>
-    if null eqns
-    then error ("primcorec error:\n  " ^ str)
-    else error ("primcorec error:\n  " ^ str ^ "\nin\n  " ^
-      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
-
-val add_primcorecursive_cmd = (the o fst) ooo add_primcorec_ursive_cmd false;
-val add_primcorec_cmd = (the o snd) ooo add_primcorec_ursive_cmd true;
-
-end;
--- a/src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML	Mon Nov 11 17:34:44 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,116 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2013
-
-Tactics for recursor and corecursor sugar.
-*)
-
-signature BNF_FP_REC_SUGAR_TACTICS =
-sig
-  val mk_primcorec_assumption_tac: Proof.context -> thm list -> int -> tactic
-  val mk_primcorec_code_of_raw_code_tac: thm list -> thm -> tactic
-  val mk_primcorec_ctr_of_dtr_tac: Proof.context -> int -> thm -> thm option -> thm list -> tactic
-  val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list ->
-    tactic
-  val mk_primcorec_raw_code_of_ctr_tac: Proof.context -> thm list -> thm list -> thm list ->
-    thm list -> int list -> thm list -> tactic
-  val mk_primcorec_sel_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
-    thm list -> thm list -> thm list -> thm -> int -> int -> thm list list list -> tactic
-  val mk_primrec_tac: Proof.context -> int -> thm list -> thm list -> thm list -> thm -> tactic
-end;
-
-structure BNF_FP_Rec_Sugar_Tactics : BNF_FP_REC_SUGAR_TACTICS =
-struct
-
-open BNF_Util
-open BNF_Tactics
-
-val falseEs = @{thms not_TrueE FalseE};
-val neq_eq_eq_contradict = @{thm neq_eq_eq_contradict};
-val split_if = @{thm split_if};
-val split_if_asm = @{thm split_if_asm};
-val split_connectI = @{thms allI impI conjI};
-
-fun mk_primrec_tac ctxt num_extra_args map_idents map_comps fun_defs recx =
-  unfold_thms_tac ctxt fun_defs THEN
-  HEADGOAL (rtac (funpow num_extra_args (fn thm => thm RS fun_cong) recx RS trans)) THEN
-  unfold_thms_tac ctxt (@{thms id_def split o_def fst_conv snd_conv} @ map_comps @ map_idents) THEN
-  HEADGOAL (rtac refl);
-
-fun mk_primcorec_assumption_tac ctxt discIs =
-  SELECT_GOAL (unfold_thms_tac ctxt
-      @{thms not_not not_False_eq_True de_Morgan_conj de_Morgan_disj} THEN
-    SOLVE (HEADGOAL (REPEAT o (rtac refl ORELSE' atac ORELSE' etac conjE ORELSE'
-    resolve_tac @{thms TrueI conjI disjI1 disjI2} ORELSE'
-    dresolve_tac discIs THEN' atac ORELSE'
-    etac notE THEN' atac ORELSE'
-    etac disjE))));
-
-fun mk_primcorec_same_case_tac m =
-  HEADGOAL (if m = 0 then rtac TrueI
-    else REPEAT_DETERM_N (m - 1) o (rtac conjI THEN' atac) THEN' atac);
-
-fun mk_primcorec_different_case_tac ctxt excl =
-  unfold_thms_tac ctxt @{thms not_not not_False_eq_True not_True_eq_False} THEN
-  HEADGOAL (rtac excl THEN_ALL_NEW mk_primcorec_assumption_tac ctxt []);
-
-fun mk_primcorec_cases_tac ctxt k m exclsss =
-  let val n = length exclsss in
-    EVERY (map (fn [] => if k = n then all_tac else mk_primcorec_same_case_tac m
-        | [excl] => mk_primcorec_different_case_tac ctxt excl)
-      (take k (nth exclsss (k - 1))))
-  end;
-
-fun mk_primcorec_prelude ctxt defs thm =
-  unfold_thms_tac ctxt defs THEN HEADGOAL (rtac thm) THEN
-  unfold_thms_tac ctxt @{thms Let_def split};
-
-fun mk_primcorec_disc_tac ctxt defs disc_corec k m exclsss =
-  mk_primcorec_prelude ctxt defs disc_corec THEN mk_primcorec_cases_tac ctxt k m exclsss;
-
-fun mk_primcorec_sel_tac ctxt defs distincts splits split_asms maps map_idents map_comps f_sel k m
-    exclsss =
-  mk_primcorec_prelude ctxt defs (f_sel RS trans) THEN
-  mk_primcorec_cases_tac ctxt k m exclsss THEN
-  HEADGOAL (REPEAT_DETERM o (rtac refl ORELSE' rtac ext ORELSE'
-    eresolve_tac falseEs ORELSE'
-    resolve_tac split_connectI ORELSE'
-    Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
-    Splitter.split_tac (split_if :: splits) ORELSE'
-    eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
-    etac notE THEN' atac ORELSE'
-    (CHANGED o SELECT_GOAL (unfold_thms_tac ctxt
-      (@{thms id_apply o_def split_def sum.cases} @ maps @ map_comps @ map_idents)))));
-
-fun mk_primcorec_ctr_of_dtr_tac ctxt m collapse maybe_disc_f sel_fs =
-  HEADGOAL (rtac ((if null sel_fs then collapse else collapse RS sym) RS trans) THEN'
-    (the_default (K all_tac) (Option.map rtac maybe_disc_f)) THEN' REPEAT_DETERM_N m o atac) THEN
-  unfold_thms_tac ctxt sel_fs THEN HEADGOAL (rtac refl);
-
-(* TODO: reduce code duplication with selector tactic above *)
-fun mk_primcorec_raw_code_of_ctr_single_tac ctxt distincts discIs splits split_asms m f_ctr =
-  HEADGOAL (REPEAT o (resolve_tac split_connectI ORELSE' split_tac (split_if :: splits))) THEN
-  mk_primcorec_prelude ctxt [] (f_ctr RS trans) THEN
-  HEADGOAL ((REPEAT_DETERM_N m o mk_primcorec_assumption_tac ctxt discIs) THEN'
-    SELECT_GOAL (SOLVE (HEADGOAL (REPEAT_DETERM o
-    (rtac refl ORELSE' atac ORELSE'
-     resolve_tac split_connectI ORELSE'
-     Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
-     Splitter.split_tac (split_if :: splits) ORELSE'
-     mk_primcorec_assumption_tac ctxt discIs ORELSE'
-     eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
-     (TRY o dresolve_tac discIs) THEN' etac notE THEN' atac)))));
-
-fun mk_primcorec_raw_code_of_ctr_tac ctxt distincts discIs splits split_asms ms ctr_thms =
-  EVERY (map2 (mk_primcorec_raw_code_of_ctr_single_tac ctxt distincts discIs splits split_asms)
-    ms ctr_thms);
-
-fun mk_primcorec_code_of_raw_code_tac splits raw =
-  HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN' REPEAT_DETERM o
-    (rtac refl ORELSE'
-     (TRY o rtac sym) THEN' atac ORELSE'
-     resolve_tac split_connectI ORELSE'
-     Splitter.split_tac (split_if :: splits) ORELSE'
-     etac notE THEN' atac));
-
-end;
--- a/src/HOL/BNF/Tools/bnf_fp_rec_sugar_util.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Tools/bnf_fp_rec_sugar_util.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -8,410 +8,26 @@
 
 signature BNF_FP_REC_SUGAR_UTIL =
 sig
-  datatype rec_call =
-    No_Rec of int |
-    Direct_Rec of int (*before*) * int (*after*) |
-    Indirect_Rec of int
-
-  datatype corec_call =
-    Dummy_No_Corec of int |
-    No_Corec of int |
-    Direct_Corec of int (*stop?*) * int (*end*) * int (*continue*) |
-    Indirect_Corec of int
-
-  type rec_ctr_spec =
-    {ctr: term,
-     offset: int,
-     calls: rec_call list,
-     rec_thm: thm}
-
-  type corec_ctr_spec =
-    {ctr: term,
-     disc: term,
-     sels: term list,
-     pred: int option,
-     calls: corec_call list,
-     discI: thm,
-     sel_thms: thm list,
-     collapse: thm,
-     corec_thm: thm,
-     disc_corec: thm,
-     sel_corecs: thm list}
+  val indexed: 'a list -> int -> int list * int
+  val indexedd: 'a list list -> int -> int list list * int
+  val indexeddd: ''a list list list -> int -> int list list list * int
+  val indexedddd: 'a list list list list -> int -> int list list list list * int
+  val find_index_eq: ''a list -> ''a -> int
+  val finds: ('a * 'b -> bool) -> 'a list -> 'b list -> ('a * 'b list) list * 'b list
 
-  type rec_spec =
-    {recx: term,
-     nested_map_idents: thm list,
-     nested_map_comps: thm list,
-     ctr_specs: rec_ctr_spec list}
-
-  type corec_spec =
-    {corec: term,
-     nested_maps: thm list,
-     nested_map_idents: thm list,
-     nested_map_comps: thm list,
-     ctr_specs: corec_ctr_spec list}
-
-  val s_not: term -> term
-  val mk_conjs: term list -> term
-  val mk_disjs: term list -> term
-  val s_not_disj: term -> term list
-  val negate_conj: term list -> term list
-  val negate_disj: term list -> term list
+  val drop_All: term -> term
 
-  val massage_indirect_rec_call: Proof.context -> (term -> bool) -> (typ -> typ -> term -> term) ->
-    typ list -> term -> term -> term -> term
-  val massage_direct_corec_call: Proof.context -> (term -> bool) -> (typ list -> term -> term) ->
-    typ list -> term -> term
-  val massage_indirect_corec_call: Proof.context -> (term -> bool) ->
-    (typ list -> typ -> typ -> term -> term) -> typ list -> typ -> term -> term
-  val expand_corec_code_rhs: Proof.context -> (term -> bool) -> typ list -> term -> term
-  val massage_corec_code_rhs: Proof.context -> (typ list -> term -> term list -> term) ->
-    typ list -> term -> term
-  val fold_rev_corec_code_rhs: Proof.context -> (term list -> term -> term list -> 'a -> 'a) ->
-    typ list -> term -> 'a -> 'a
-  val case_thms_of_term: Proof.context -> typ list -> term ->
-    thm list * thm list * thm list * thm list
+  val mk_partial_compN: int -> typ -> term -> term
+  val mk_partial_comp: typ -> typ -> term -> term
+  val mk_compN: int -> typ list -> term * term -> term
+  val mk_comp: typ list -> term * term -> term
 
-  val rec_specs_of: binding list -> typ list -> typ list -> (term -> int list) ->
-    ((term * term list list) list) list -> local_theory ->
-    (bool * rec_spec list * typ list * thm * thm list) * local_theory
-  val corec_specs_of: binding list -> typ list -> typ list -> (term -> int list) ->
-    ((term * term list list) list) list -> local_theory ->
-    (bool * corec_spec list * typ list * thm * thm * thm list * thm list) * local_theory
+  val get_indices: ((binding * typ) * 'a) list -> term -> int list
 end;
 
 structure BNF_FP_Rec_Sugar_Util : BNF_FP_REC_SUGAR_UTIL =
 struct
 
-open Ctr_Sugar
-open BNF_Util
-open BNF_Def
-open BNF_FP_Util
-open BNF_FP_Def_Sugar
-open BNF_FP_N2M_Sugar
-
-datatype rec_call =
-  No_Rec of int |
-  Direct_Rec of int * int |
-  Indirect_Rec of int;
-
-datatype corec_call =
-  Dummy_No_Corec of int |
-  No_Corec of int |
-  Direct_Corec of int * int * int |
-  Indirect_Corec of int;
-
-type rec_ctr_spec =
-  {ctr: term,
-   offset: int,
-   calls: rec_call list,
-   rec_thm: thm};
-
-type corec_ctr_spec =
-  {ctr: term,
-   disc: term,
-   sels: term list,
-   pred: int option,
-   calls: corec_call list,
-   discI: thm,
-   sel_thms: thm list,
-   collapse: thm,
-   corec_thm: thm,
-   disc_corec: thm,
-   sel_corecs: thm list};
-
-type rec_spec =
-  {recx: term,
-   nested_map_idents: thm list,
-   nested_map_comps: thm list,
-   ctr_specs: rec_ctr_spec list};
-
-type corec_spec =
-  {corec: term,
-   nested_maps: thm list,
-   nested_map_idents: thm list,
-   nested_map_comps: thm list,
-   ctr_specs: corec_ctr_spec list};
-
-val id_def = @{thm id_def};
-
-exception AINT_NO_MAP of term;
-
-fun ill_formed_rec_call ctxt t =
-  error ("Ill-formed recursive call: " ^ quote (Syntax.string_of_term ctxt t));
-fun ill_formed_corec_call ctxt t =
-  error ("Ill-formed corecursive call: " ^ quote (Syntax.string_of_term ctxt t));
-fun invalid_map ctxt t =
-  error ("Invalid map function in " ^ quote (Syntax.string_of_term ctxt t));
-fun unexpected_rec_call ctxt t =
-  error ("Unexpected recursive call: " ^ quote (Syntax.string_of_term ctxt t));
-fun unexpected_corec_call ctxt t =
-  error ("Unexpected corecursive call: " ^ quote (Syntax.string_of_term ctxt t));
-
-fun s_not @{const True} = @{const False}
-  | s_not @{const False} = @{const True}
-  | s_not (@{const Not} $ t) = t
-  | s_not t = HOLogic.mk_not t
-
-val mk_conjs = try (foldr1 HOLogic.mk_conj) #> the_default @{const True};
-val mk_disjs = try (foldr1 HOLogic.mk_disj) #> the_default @{const False};
-
-val s_not_disj = map s_not o HOLogic.disjuncts;
-
-fun negate_conj [t] = s_not_disj t
-  | negate_conj ts = [mk_disjs (map s_not ts)];
-
-fun negate_disj [t] = s_not_disj t
-  | negate_disj ts = [mk_disjs (map (mk_conjs o s_not_disj) ts)];
-
-fun factor_out_types ctxt massage destU U T =
-  (case try destU U of
-    SOME (U1, U2) => if U1 = T then massage T U2 else invalid_map ctxt
-  | NONE => invalid_map ctxt);
-
-fun map_flattened_map_args ctxt s map_args fs =
-  let
-    val flat_fs = flatten_type_args_of_bnf (the (bnf_of ctxt s)) Term.dummy fs;
-    val flat_fs' = map_args flat_fs;
-  in
-    permute_like (op aconv) flat_fs fs flat_fs'
-  end;
-
-fun massage_indirect_rec_call ctxt has_call raw_massage_fun bound_Ts y y' =
-  let
-    val typof = curry fastype_of1 bound_Ts;
-    val build_map_fst = build_map ctxt (fst_const o fst);
-
-    val yT = typof y;
-    val yU = typof y';
-
-    fun y_of_y' () = build_map_fst (yU, yT) $ y';
-    val elim_y = Term.map_aterms (fn t => if t = y then y_of_y' () else t);
-
-    fun massage_direct_fun U T t =
-      if has_call t then factor_out_types ctxt raw_massage_fun HOLogic.dest_prodT U T t
-      else HOLogic.mk_comp (t, build_map_fst (U, T));
-
-    fun massage_map (Type (_, Us)) (Type (s, Ts)) t =
-        (case try (dest_map ctxt s) t of
-          SOME (map0, fs) =>
-          let
-            val Type (_, ran_Ts) = range_type (typof t);
-            val map' = mk_map (length fs) Us ran_Ts map0;
-            val fs' = map_flattened_map_args ctxt s (map3 massage_map_or_map_arg Us Ts) fs;
-          in
-            Term.list_comb (map', fs')
-          end
-        | NONE => raise AINT_NO_MAP t)
-      | massage_map _ _ t = raise AINT_NO_MAP t
-    and massage_map_or_map_arg U T t =
-      if T = U then
-        if has_call t then unexpected_rec_call ctxt t else t
-      else
-        massage_map U T t
-        handle AINT_NO_MAP _ => massage_direct_fun U T t;
-
-    fun massage_call (t as t1 $ t2) =
-        if t2 = y then
-          massage_map yU yT (elim_y t1) $ y'
-          handle AINT_NO_MAP t' => invalid_map ctxt t'
-        else
-          ill_formed_rec_call ctxt t
-      | massage_call t = if t = y then y_of_y' () else ill_formed_rec_call ctxt t;
-  in
-    massage_call
-  end;
-
-fun fold_rev_let_if_case ctxt f bound_Ts t =
-  let
-    val thy = Proof_Context.theory_of ctxt;
-
-    fun fld conds t =
-      (case Term.strip_comb t of
-        (Const (@{const_name Let}, _), [arg1, arg2]) => fld conds (betapply (arg2, arg1))
-      | (Const (@{const_name If}, _), [cond, then_branch, else_branch]) =>
-        fld (conds @ HOLogic.conjuncts cond) then_branch
-        o fld (conds @ s_not_disj cond) else_branch
-      | (Const (c, _), args as _ :: _ :: _) =>
-        let val n = num_binder_types (Sign.the_const_type thy c) - 1 in
-          if n >= 0 andalso n < length args then
-            (case fastype_of1 (bound_Ts, nth args n) of
-              Type (s, Ts) =>
-              (case dest_case ctxt s Ts t of
-                NONE => apsnd (f conds t)
-              | SOME (conds', branches) =>
-                apfst (cons s) o fold_rev (uncurry fld)
-                  (map (append conds o HOLogic.conjuncts) conds' ~~ branches))
-            | _ => apsnd (f conds t))
-          else
-            apsnd (f conds t)
-        end
-      | _ => apsnd (f conds t))
-  in
-    fld [] t o pair []
-  end;
-
-fun case_of ctxt = ctr_sugar_of ctxt #> Option.map (fst o dest_Const o #casex);
-
-fun massage_let_if_case ctxt has_call massage_leaf =
-  let
-    val thy = Proof_Context.theory_of ctxt;
-
-    fun check_no_call t = if has_call t then unexpected_corec_call ctxt t else ();
-
-    fun massage_abs bound_Ts (Abs (s, T, t)) = Abs (s, T, massage_abs (T :: bound_Ts) t)
-      | massage_abs bound_Ts t = massage_rec bound_Ts t
-    and massage_rec bound_Ts t =
-      let val typof = curry fastype_of1 bound_Ts in
-        (case Term.strip_comb t of
-          (Const (@{const_name Let}, _), [arg1, arg2]) =>
-          massage_rec bound_Ts (betapply (arg2, arg1))
-        | (Const (@{const_name If}, _), obj :: (branches as [_, _])) =>
-          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
-        | (Const (c, _), args as _ :: _ :: _) =>
-          let
-            val gen_T = Sign.the_const_type thy c;
-            val (gen_branch_Ts, gen_body_fun_T) = strip_fun_type gen_T;
-            val n = length gen_branch_Ts;
-          in
-            if n < length args then
-              (case gen_body_fun_T of
-                Type (_, [Type (T_name, _), _]) =>
-                if case_of ctxt T_name = SOME c then
-                  let
-                    val (branches, obj_leftovers) = chop n args;
-                    val branches' = map (massage_abs bound_Ts o Envir.eta_long bound_Ts) branches;
-                    val branch_Ts' = map typof branches';
-                    val casex' = Const (c, branch_Ts' ---> map typof obj_leftovers --->
-                      snd (strip_typeN (num_binder_types (hd gen_branch_Ts)) (hd branch_Ts')));
-                  in
-                    Term.list_comb (casex', branches' @ tap (List.app check_no_call) obj_leftovers)
-                  end
-                else
-                  massage_leaf bound_Ts t
-              | _ => massage_leaf bound_Ts t)
-            else
-              massage_leaf bound_Ts t
-          end
-        | _ => massage_leaf bound_Ts t)
-      end
-  in
-    massage_rec
-  end;
-
-val massage_direct_corec_call = massage_let_if_case;
-
-fun curried_type (Type (@{type_name fun}, [Type (@{type_name prod}, Ts), T])) = Ts ---> T;
-
-fun massage_indirect_corec_call ctxt has_call raw_massage_call bound_Ts U t =
-  let
-    val build_map_Inl = build_map ctxt (uncurry Inl_const o dest_sumT o snd)
-
-    fun massage_direct_call bound_Ts U T t =
-      if has_call t then factor_out_types ctxt (raw_massage_call bound_Ts) dest_sumT U T t
-      else build_map_Inl (T, U) $ t;
-
-    fun massage_direct_fun bound_Ts U T t =
-      let
-        val var = Var ((Name.uu, Term.maxidx_of_term t + 1),
-          domain_type (fastype_of1 (bound_Ts, t)));
-      in
-        Term.lambda var (massage_direct_call bound_Ts U T (t $ var))
-      end;
-
-    fun massage_map bound_Ts (Type (_, Us)) (Type (s, Ts)) t =
-        (case try (dest_map ctxt s) t of
-          SOME (map0, fs) =>
-          let
-            val Type (_, dom_Ts) = domain_type (fastype_of1 (bound_Ts, t));
-            val map' = mk_map (length fs) dom_Ts Us map0;
-            val fs' =
-              map_flattened_map_args ctxt s (map3 (massage_map_or_map_arg bound_Ts) Us Ts) fs;
-          in
-            Term.list_comb (map', fs')
-          end
-        | NONE => raise AINT_NO_MAP t)
-      | massage_map _ _ _ t = raise AINT_NO_MAP t
-    and massage_map_or_map_arg bound_Ts U T t =
-      if T = U then
-        if has_call t then unexpected_corec_call ctxt t else t
-      else
-        massage_map bound_Ts U T t
-        handle AINT_NO_MAP _ => massage_direct_fun bound_Ts U T t;
-
-    fun massage_call bound_Ts U T =
-      massage_let_if_case ctxt has_call (fn bound_Ts => fn t =>
-        if has_call t then
-          (case U of
-            Type (s, Us) =>
-            (case try (dest_ctr ctxt s) t of
-              SOME (f, args) =>
-              let
-                val typof = curry fastype_of1 bound_Ts;
-                val f' = mk_ctr Us f
-                val f'_T = typof f';
-                val arg_Ts = map typof args;
-              in
-                Term.list_comb (f', map3 (massage_call bound_Ts) (binder_types f'_T) arg_Ts args)
-              end
-            | NONE =>
-              (case t of
-                Const (@{const_name prod_case}, _) $ 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'
-                end
-              | t1 $ t2 =>
-                (if has_call t2 then
-                  massage_direct_call bound_Ts U T t
-                else
-                  massage_map bound_Ts U T t1 $ t2
-                  handle AINT_NO_MAP _ => massage_direct_call bound_Ts U T t)
-              | Abs (s, T', t') =>
-                Abs (s, T', massage_call (T' :: bound_Ts) (range_type U) (range_type T) t')
-              | _ => massage_direct_call bound_Ts U T t))
-          | _ => ill_formed_corec_call ctxt t)
-        else
-          build_map_Inl (T, U) $ t) bound_Ts;
-
-    val T = fastype_of1 (bound_Ts, t);
-  in
-    if has_call t then massage_call bound_Ts U T t else build_map_Inl (T, U) $ t
-  end;
-
-fun expand_ctr_term ctxt s Ts t =
-  (case ctr_sugar_of ctxt s of
-    SOME {ctrs, casex, ...} =>
-    Term.list_comb (mk_case Ts (Type (s, Ts)) casex, map (mk_ctr Ts) ctrs) $ t
-  | NONE => raise Fail "expand_ctr_term");
-
-fun expand_corec_code_rhs ctxt has_call bound_Ts t =
-  (case fastype_of1 (bound_Ts, t) of
-    Type (s, Ts) =>
-    massage_let_if_case ctxt has_call (fn _ => fn t =>
-      if can (dest_ctr ctxt s) t then t else expand_ctr_term ctxt s Ts t) bound_Ts t
-  | _ => raise Fail "expand_corec_code_rhs");
-
-fun massage_corec_code_rhs ctxt massage_ctr =
-  massage_let_if_case ctxt (K false)
-    (fn bound_Ts => uncurry (massage_ctr bound_Ts) o Term.strip_comb);
-
-fun fold_rev_corec_code_rhs ctxt f =
-  snd ooo 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
-    val (caseT_names, _) = fold_rev_let_if_case ctxt (K (K I)) bound_Ts t ();
-    val ctr_sugars = map (the o ctr_sugar_of ctxt) caseT_names;
-  in
-    (maps #distincts ctr_sugars, maps #discIs ctr_sugars, maps #sel_splits ctr_sugars,
-     maps #sel_split_asms ctr_sugars)
-  end;
-
 fun indexed xs h = let val h' = h + length xs in (h upto h' - 1, h') end;
 fun indexedd xss = fold_map indexed xss;
 fun indexeddd xsss = fold_map indexedd xsss;
@@ -419,205 +35,32 @@
 
 fun find_index_eq hs h = find_index (curry (op =) h) hs;
 
-(*FIXME: remove special cases for products and sum once they are registered as datatypes*)
-fun map_thms_of_typ ctxt (Type (s, _)) =
-    if s = @{type_name prod} then
-      @{thms map_pair_simp}
-    else if s = @{type_name sum} then
-      @{thms sum_map.simps}
-    else
-      (case fp_sugar_of ctxt s of
-        SOME {index, mapss, ...} => nth mapss index
-      | NONE => [])
-  | map_thms_of_typ _ _ = [];
-
-fun rec_specs_of bs arg_Ts res_Ts get_indices callssss0 lthy =
-  let
-    val thy = Proof_Context.theory_of lthy;
-
-    val ((missing_arg_Ts, perm0_kks,
-          fp_sugars as {nested_bnfs, fp_res = {xtor_co_iterss = ctor_iters1 :: _, ...},
-            co_inducts = [induct_thm], ...} :: _, (lfp_sugar_thms, _)), lthy') =
-      nested_to_mutual_fps Least_FP bs arg_Ts get_indices callssss0 lthy;
-
-    val perm_fp_sugars = sort (int_ord o pairself #index) fp_sugars;
-
-    val indices = map #index fp_sugars;
-    val perm_indices = map #index perm_fp_sugars;
-
-    val perm_ctrss = map (#ctrs o of_fp_sugar #ctr_sugars) perm_fp_sugars;
-    val perm_ctr_Tsss = map (map (binder_types o fastype_of)) perm_ctrss;
-    val perm_fpTs = map (body_type o fastype_of o hd) perm_ctrss;
-
-    val nn0 = length arg_Ts;
-    val nn = length perm_fpTs;
-    val kks = 0 upto nn - 1;
-    val perm_ns = map length perm_ctr_Tsss;
-    val perm_mss = map (map length) perm_ctr_Tsss;
-
-    val perm_Cs = map (body_type o fastype_of o co_rec_of o of_fp_sugar (#xtor_co_iterss o #fp_res))
-      perm_fp_sugars;
-    val perm_fun_arg_Tssss =
-      mk_iter_fun_arg_types perm_ctr_Tsss perm_ns perm_mss (co_rec_of ctor_iters1);
-
-    fun unpermute0 perm0_xs = permute_like (op =) perm0_kks kks perm0_xs;
-    fun unpermute perm_xs = permute_like (op =) perm_indices indices perm_xs;
-
-    val induct_thms = unpermute0 (conj_dests nn induct_thm);
+fun finds eq = fold_map (fn x => List.partition (curry eq x) #>> pair x);
 
-    val fpTs = unpermute perm_fpTs;
-    val Cs = unpermute perm_Cs;
-
-    val As_rho = tvar_subst thy (take nn0 fpTs) arg_Ts;
-    val Cs_rho = map (fst o dest_TVar) Cs ~~ pad_list HOLogic.unitT nn res_Ts;
-
-    val substA = Term.subst_TVars As_rho;
-    val substAT = Term.typ_subst_TVars As_rho;
-    val substCT = Term.typ_subst_TVars Cs_rho;
-
-    val perm_Cs' = map substCT perm_Cs;
-
-    fun offset_of_ctr 0 _ = 0
-      | offset_of_ctr n (({ctrs, ...} : ctr_sugar) :: ctr_sugars) =
-        length ctrs + offset_of_ctr (n - 1) ctr_sugars;
-
-    fun call_of [i] [T] = (if exists_subtype_in Cs T then Indirect_Rec else No_Rec) i
-      | call_of [i, i'] _ = Direct_Rec (i, i');
+fun drop_All t =
+  subst_bounds (strip_qnt_vars @{const_name all} t |> map Free |> rev,
+    strip_qnt_body @{const_name all} t);
 
-    fun mk_ctr_spec ctr offset fun_arg_Tss rec_thm =
-      let
-        val (fun_arg_hss, _) = indexedd fun_arg_Tss 0;
-        val fun_arg_hs = flat_rec_arg_args fun_arg_hss;
-        val fun_arg_iss = map (map (find_index_eq fun_arg_hs)) fun_arg_hss;
-      in
-        {ctr = substA ctr, offset = offset, calls = map2 call_of fun_arg_iss fun_arg_Tss,
-         rec_thm = rec_thm}
-      end;
-
-    fun mk_ctr_specs index (ctr_sugars : ctr_sugar list) iter_thmsss =
-      let
-        val ctrs = #ctrs (nth ctr_sugars index);
-        val rec_thmss = co_rec_of (nth iter_thmsss index);
-        val k = offset_of_ctr index ctr_sugars;
-        val n = length ctrs;
-      in
-        map4 mk_ctr_spec ctrs (k upto k + n - 1) (nth perm_fun_arg_Tssss index) rec_thmss
-      end;
-
-    fun mk_spec ({T, index, ctr_sugars, co_iterss = iterss, co_iter_thmsss = iter_thmsss, ...}
-      : fp_sugar) =
-      {recx = mk_co_iter thy Least_FP (substAT T) perm_Cs' (co_rec_of (nth iterss index)),
-       nested_map_idents = map (unfold_thms lthy [id_def] o map_id0_of_bnf) nested_bnfs,
-       nested_map_comps = map map_comp_of_bnf nested_bnfs,
-       ctr_specs = mk_ctr_specs index ctr_sugars iter_thmsss};
-  in
-    ((is_some lfp_sugar_thms, map mk_spec fp_sugars, missing_arg_Ts, induct_thm, induct_thms),
-     lthy')
+fun mk_partial_comp gT fT g =
+  let val T = domain_type fT --> range_type gT in
+    Const (@{const_name Fun.comp}, gT --> fT --> T) $ g
   end;
 
-fun corec_specs_of bs arg_Ts res_Ts get_indices callssss0 lthy =
-  let
-    val thy = Proof_Context.theory_of lthy;
-
-    val ((missing_res_Ts, perm0_kks,
-          fp_sugars as {nested_bnfs, fp_res = {xtor_co_iterss = dtor_coiters1 :: _, ...},
-            co_inducts = coinduct_thms, ...} :: _, (_, gfp_sugar_thms)), lthy') =
-      nested_to_mutual_fps Greatest_FP bs res_Ts get_indices callssss0 lthy;
-
-    val perm_fp_sugars = sort (int_ord o pairself #index) fp_sugars;
-
-    val indices = map #index fp_sugars;
-    val perm_indices = map #index perm_fp_sugars;
-
-    val perm_ctrss = map (#ctrs o of_fp_sugar #ctr_sugars) perm_fp_sugars;
-    val perm_ctr_Tsss = map (map (binder_types o fastype_of)) perm_ctrss;
-    val perm_fpTs = map (body_type o fastype_of o hd) perm_ctrss;
-
-    val nn0 = length res_Ts;
-    val nn = length perm_fpTs;
-    val kks = 0 upto nn - 1;
-    val perm_ns = map length perm_ctr_Tsss;
-
-    val perm_Cs = map (domain_type o body_fun_type o fastype_of o co_rec_of o
-      of_fp_sugar (#xtor_co_iterss o #fp_res)) perm_fp_sugars;
-    val (perm_p_Tss, (perm_q_Tssss, _, perm_f_Tssss, _)) =
-      mk_coiter_fun_arg_types perm_ctr_Tsss perm_Cs perm_ns (co_rec_of dtor_coiters1);
-
-    val (perm_p_hss, h) = indexedd perm_p_Tss 0;
-    val (perm_q_hssss, h') = indexedddd perm_q_Tssss h;
-    val (perm_f_hssss, _) = indexedddd perm_f_Tssss h';
-
-    val fun_arg_hs =
-      flat (map3 flat_corec_preds_predsss_gettersss perm_p_hss perm_q_hssss perm_f_hssss);
-
-    fun unpermute0 perm0_xs = permute_like (op =) perm0_kks kks perm0_xs;
-    fun unpermute perm_xs = permute_like (op =) perm_indices indices perm_xs;
-
-    val coinduct_thmss = map (unpermute0 o conj_dests nn) coinduct_thms;
-
-    val p_iss = map (map (find_index_eq fun_arg_hs)) (unpermute perm_p_hss);
-    val q_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_q_hssss);
-    val f_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_f_hssss);
-
-    val f_Tssss = unpermute perm_f_Tssss;
-    val fpTs = unpermute perm_fpTs;
-    val Cs = unpermute perm_Cs;
-
-    val As_rho = tvar_subst thy (take nn0 fpTs) res_Ts;
-    val Cs_rho = map (fst o dest_TVar) Cs ~~ pad_list HOLogic.unitT nn arg_Ts;
+fun mk_partial_compN 0 _ g = g
+  | mk_partial_compN n fT g =
+    let val g' = mk_partial_compN (n - 1) (range_type fT) g in
+      mk_partial_comp (fastype_of g') fT g'
+    end;
 
-    val substA = Term.subst_TVars As_rho;
-    val substAT = Term.typ_subst_TVars As_rho;
-    val substCT = Term.typ_subst_TVars Cs_rho;
-
-    val perm_Cs' = map substCT perm_Cs;
-
-    fun call_of nullary [] [g_i] [Type (@{type_name fun}, [_, T])] =
-        (if exists_subtype_in Cs T then Indirect_Corec
-         else if nullary then Dummy_No_Corec
-         else No_Corec) g_i
-      | call_of _ [q_i] [g_i, g_i'] _ = Direct_Corec (q_i, g_i, g_i');
-
-    fun mk_ctr_spec ctr disc sels p_ho q_iss f_iss f_Tss discI sel_thms collapse corec_thm
-        disc_corec sel_corecs =
-      let val nullary = not (can dest_funT (fastype_of ctr)) in
-        {ctr = substA ctr, disc = substA disc, sels = map substA sels, pred = p_ho,
-         calls = map3 (call_of nullary) q_iss f_iss f_Tss, discI = discI, sel_thms = sel_thms,
-         collapse = collapse, corec_thm = corec_thm, disc_corec = disc_corec,
-         sel_corecs = sel_corecs}
-      end;
-
-    fun mk_ctr_specs index (ctr_sugars : ctr_sugar list) p_is q_isss f_isss f_Tsss
-        coiter_thmsss disc_coitersss sel_coiterssss =
-      let
-        val ctrs = #ctrs (nth ctr_sugars index);
-        val discs = #discs (nth ctr_sugars index);
-        val selss = #selss (nth ctr_sugars index);
-        val p_ios = map SOME p_is @ [NONE];
-        val discIs = #discIs (nth ctr_sugars index);
-        val sel_thmss = #sel_thmss (nth ctr_sugars index);
-        val collapses = #collapses (nth ctr_sugars index);
-        val corec_thms = co_rec_of (nth coiter_thmsss index);
-        val disc_corecs = co_rec_of (nth disc_coitersss index);
-        val sel_corecss = co_rec_of (nth sel_coiterssss index);
-      in
-        map13 mk_ctr_spec ctrs discs selss p_ios q_isss f_isss f_Tsss discIs sel_thmss collapses
-          corec_thms disc_corecs sel_corecss
-      end;
-
-    fun mk_spec ({T, index, ctr_sugars, co_iterss = coiterss, co_iter_thmsss = coiter_thmsss,
-          disc_co_itersss = disc_coitersss, sel_co_iterssss = sel_coiterssss, ...} : fp_sugar)
-        p_is q_isss f_isss f_Tsss =
-      {corec = mk_co_iter thy Greatest_FP (substAT T) perm_Cs' (co_rec_of (nth coiterss index)),
-       nested_maps = maps (map_thms_of_typ lthy o T_of_bnf) nested_bnfs,
-       nested_map_idents = map (unfold_thms lthy [id_def] o map_id0_of_bnf) nested_bnfs,
-       nested_map_comps = map map_comp_of_bnf nested_bnfs,
-       ctr_specs = mk_ctr_specs index ctr_sugars p_is q_isss f_isss f_Tsss coiter_thmsss
-         disc_coitersss sel_coiterssss};
-  in
-    ((is_some gfp_sugar_thms, map5 mk_spec fp_sugars p_iss q_issss f_issss f_Tssss, missing_res_Ts,
-      co_induct_of coinduct_thms, strong_co_induct_of coinduct_thms, co_induct_of coinduct_thmss,
-      strong_co_induct_of coinduct_thmss), lthy')
+fun mk_compN n bound_Ts (g, f) =
+  let val typof = curry fastype_of1 bound_Ts in
+    mk_partial_compN n (typof f) g $ f
   end;
 
+val mk_comp = mk_compN 1;
+
+fun get_indices fixes t = map (fst #>> Binding.name_of #> Free) fixes
+  |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
+  |> map_filter I;
+
 end;
--- a/src/HOL/BNF/Tools/bnf_gfp.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Tools/bnf_gfp.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -23,7 +23,7 @@
 open BNF_Comp
 open BNF_FP_Util
 open BNF_FP_Def_Sugar
-open BNF_FP_Rec_Sugar
+open BNF_GFP_Rec_Sugar
 open BNF_GFP_Util
 open BNF_GFP_Tactics
 
@@ -2744,8 +2744,8 @@
               ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
             bs thmss)
       in
-       (timer, Jbnfs, (folded_dtor_map_o_thms, folded_dtor_map_thms), folded_dtor_set_thmss',
-         dtor_set_induct_thms, dtor_Jrel_thms, Jbnf_common_notes @ Jbnf_notes, lthy)
+        (timer, Jbnfs, (folded_dtor_map_o_thms, folded_dtor_map_thms), folded_dtor_set_thmss',
+          dtor_set_induct_thms, dtor_Jrel_thms, Jbnf_common_notes @ Jbnf_notes, lthy)
       end;
 
       val dtor_unfold_o_map_thms = mk_xtor_un_fold_o_map_thms Greatest_FP false m
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_gfp_rec_sugar.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -0,0 +1,1150 @@
+(*  Title:      HOL/BNF/Tools/bnf_gfp_rec_sugar.ML
+    Author:     Lorenz Panny, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2013
+
+Corecursor sugar.
+*)
+
+signature BNF_GFP_REC_SUGAR =
+sig
+  val add_primcorecursive_cmd: bool ->
+    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
+    Proof.context -> Proof.state
+  val add_primcorec_cmd: bool ->
+    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
+    local_theory -> local_theory
+end;
+
+structure BNF_GFP_Rec_Sugar : BNF_GFP_REC_SUGAR =
+struct
+
+open Ctr_Sugar
+open BNF_Util
+open BNF_Def
+open BNF_FP_Util
+open BNF_FP_Def_Sugar
+open BNF_FP_N2M_Sugar
+open BNF_FP_Rec_Sugar_Util
+open BNF_GFP_Rec_Sugar_Tactics
+
+val codeN = "code"
+val ctrN = "ctr"
+val discN = "disc"
+val selN = "sel"
+
+val nitpicksimp_attrs = @{attributes [nitpick_simp]};
+val simp_attrs = @{attributes [simp]};
+val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
+
+exception Primcorec_Error of string * term list;
+
+fun primcorec_error str = raise Primcorec_Error (str, []);
+fun primcorec_error_eqn str eqn = raise Primcorec_Error (str, [eqn]);
+fun primcorec_error_eqns str eqns = raise Primcorec_Error (str, eqns);
+
+datatype corec_call =
+  Dummy_No_Corec of int |
+  No_Corec of int |
+  Mutual_Corec of int * int * int |
+  Nested_Corec of int;
+
+type basic_corec_ctr_spec =
+  {ctr: term,
+   disc: term,
+   sels: term list};
+
+type corec_ctr_spec =
+  {ctr: term,
+   disc: term,
+   sels: term list,
+   pred: int option,
+   calls: corec_call list,
+   discI: thm,
+   sel_thms: thm list,
+   collapse: thm,
+   corec_thm: thm,
+   disc_corec: thm,
+   sel_corecs: thm list};
+
+type corec_spec =
+  {corec: term,
+   nested_map_idents: thm list,
+   nested_map_comps: thm list,
+   ctr_specs: corec_ctr_spec list};
+
+exception AINT_NO_MAP of term;
+
+fun not_codatatype ctxt T =
+  error ("Not a codatatype: " ^ Syntax.string_of_typ ctxt T);
+fun ill_formed_corec_call ctxt t =
+  error ("Ill-formed corecursive call: " ^ quote (Syntax.string_of_term ctxt t));
+fun invalid_map ctxt t =
+  error ("Invalid map function in " ^ quote (Syntax.string_of_term ctxt t));
+fun unexpected_corec_call ctxt t =
+  error ("Unexpected corecursive call: " ^ quote (Syntax.string_of_term ctxt t));
+
+val mk_conjs = try (foldr1 HOLogic.mk_conj) #> the_default @{const True};
+val mk_disjs = try (foldr1 HOLogic.mk_disj) #> the_default @{const False};
+
+val conjuncts_s = filter_out (curry (op =) @{const True}) o HOLogic.conjuncts;
+
+fun s_not @{const True} = @{const False}
+  | s_not @{const False} = @{const True}
+  | s_not (@{const Not} $ t) = t
+  | s_not (@{const conj} $ t $ u) = @{const disj} $ s_not t $ s_not u
+  | s_not (@{const disj} $ t $ u) = @{const conj} $ s_not t $ s_not u
+  | s_not t = @{const Not} $ t;
+
+val s_not_conj = conjuncts_s o s_not o mk_conjs;
+
+fun propagate_unit_pos u cs = if member (op aconv) cs u then [@{const False}] else cs;
+
+fun propagate_unit_neg not_u cs = remove (op aconv) not_u cs;
+
+fun propagate_units css =
+  (case List.partition (can the_single) css of
+     ([], _) => css
+   | ([u] :: uss, css') =>
+     [u] :: propagate_units (map (propagate_unit_neg (s_not u))
+       (map (propagate_unit_pos u) (uss @ css'))));
+
+fun s_conjs cs =
+  if member (op aconv) cs @{const False} then @{const False}
+  else mk_conjs (remove (op aconv) @{const True} cs);
+
+fun s_disjs ds =
+  if member (op aconv) ds @{const True} then @{const True}
+  else mk_disjs (remove (op aconv) @{const False} ds);
+
+fun s_dnf css0 =
+  let val css = propagate_units css0 in
+    if null css then
+      [@{const False}]
+    else if exists null css then
+      []
+    else
+      map (fn c :: cs => (c, cs)) css
+      |> AList.coalesce (op =)
+      |> map (fn (c, css) => c :: s_dnf css)
+      |> (fn [cs] => cs | css => [s_disjs (map s_conjs css)])
+  end;
+
+fun fold_rev_let_if_case ctxt f bound_Ts t =
+  let
+    val thy = Proof_Context.theory_of ctxt;
+
+    fun fld conds t =
+      (case Term.strip_comb t of
+        (Const (@{const_name Let}, _), [_, _]) => fld conds (unfold_let t)
+      | (Const (@{const_name If}, _), [cond, then_branch, else_branch]) =>
+        fld (conds @ conjuncts_s cond) then_branch o fld (conds @ s_not_conj [cond]) else_branch
+      | (Const (c, _), args as _ :: _ :: _) =>
+        let val n = num_binder_types (Sign.the_const_type thy c) - 1 in
+          if n >= 0 andalso n < length args then
+            (case fastype_of1 (bound_Ts, nth args n) of
+              Type (s, Ts) =>
+              (case dest_case ctxt s Ts t of
+                NONE => apsnd (f conds t)
+              | SOME (conds', branches) =>
+                apfst (cons s) o fold_rev (uncurry fld)
+                  (map (append conds o conjuncts_s) conds' ~~ branches))
+            | _ => apsnd (f conds t))
+          else
+            apsnd (f conds t)
+        end
+      | _ => apsnd (f conds t))
+  in
+    fld [] t o pair []
+  end;
+
+fun case_of ctxt = ctr_sugar_of ctxt #> Option.map (fst o dest_Const o #casex);
+
+fun massage_let_if_case ctxt has_call massage_leaf =
+  let
+    val thy = Proof_Context.theory_of ctxt;
+
+    fun check_no_call t = if has_call t then unexpected_corec_call ctxt t else ();
+
+    fun massage_abs bound_Ts 0 t = massage_rec bound_Ts t
+      | massage_abs bound_Ts m (Abs (s, T, t)) = Abs (s, T, massage_abs (T :: bound_Ts) (m - 1) t)
+      | massage_abs bound_Ts m t =
+        let val T = domain_type (fastype_of1 (bound_Ts, t)) in
+          Abs (Name.uu, T, massage_abs (T :: bound_Ts) (m - 1) (incr_boundvars 1 t $ Bound 0))
+        end
+    and massage_rec bound_Ts t =
+      let val typof = curry fastype_of1 bound_Ts in
+        (case Term.strip_comb t of
+          (Const (@{const_name Let}, _), [_, _]) => massage_rec bound_Ts (unfold_let t)
+        | (Const (@{const_name If}, _), obj :: (branches as [_, _])) =>
+          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
+        | (Const (c, _), args as _ :: _ :: _) =>
+          (case try strip_fun_type (Sign.the_const_type thy c) of
+            SOME (gen_branch_Ts, gen_body_fun_T) =>
+            let
+              val gen_branch_ms = map num_binder_types gen_branch_Ts;
+              val n = length gen_branch_ms;
+            in
+              if n < length args then
+                (case gen_body_fun_T of
+                  Type (_, [Type (T_name, _), _]) =>
+                  if case_of ctxt T_name = SOME c then
+                    let
+                      val (branches, obj_leftovers) = chop n args;
+                      val branches' = map2 (massage_abs bound_Ts) gen_branch_ms branches;
+                      val branch_Ts' = map typof branches';
+                      val body_T' = snd (strip_typeN (hd gen_branch_ms) (hd branch_Ts'));
+                      val casex' = Const (c, branch_Ts' ---> map typof obj_leftovers ---> body_T');
+                    in
+                      Term.list_comb (casex',
+                        branches' @ tap (List.app check_no_call) obj_leftovers)
+                    end
+                  else
+                    massage_leaf bound_Ts t
+                | _ => massage_leaf bound_Ts t)
+              else
+                massage_leaf bound_Ts t
+            end
+          | NONE => massage_leaf bound_Ts t)
+        | _ => massage_leaf bound_Ts t)
+      end
+  in
+    massage_rec
+  end;
+
+val massage_mutual_corec_call = massage_let_if_case;
+
+fun curried_type (Type (@{type_name fun}, [Type (@{type_name prod}, Ts), T])) = Ts ---> T;
+
+fun massage_nested_corec_call ctxt has_call raw_massage_call bound_Ts U t =
+  let
+    fun check_no_call t = if has_call t then unexpected_corec_call ctxt t else ();
+
+    val build_map_Inl = build_map ctxt (uncurry Inl_const o dest_sumT o snd);
+
+    fun massage_mutual_call bound_Ts U T t =
+      if has_call t then
+        (case try dest_sumT U of
+          SOME (U1, U2) => if U1 = T then raw_massage_call bound_Ts T U2 t else invalid_map ctxt t
+        | NONE => invalid_map ctxt t)
+      else
+        build_map_Inl (T, U) $ t;
+
+    fun massage_mutual_fun bound_Ts U T t =
+      (case t of
+        Const (@{const_name comp}, _) $ t1 $ t2 =>
+        mk_comp bound_Ts (massage_mutual_fun bound_Ts U T t1, tap check_no_call t2)
+      | _ =>
+        let
+          val var = Var ((Name.uu, Term.maxidx_of_term t + 1),
+            domain_type (fastype_of1 (bound_Ts, t)));
+        in
+          Term.lambda var (massage_mutual_call bound_Ts U T (t $ var))
+        end);
+
+    fun massage_map bound_Ts (Type (_, Us)) (Type (s, Ts)) t =
+        (case try (dest_map ctxt s) t of
+          SOME (map0, fs) =>
+          let
+            val Type (_, dom_Ts) = domain_type (fastype_of1 (bound_Ts, t));
+            val map' = mk_map (length fs) dom_Ts Us map0;
+            val fs' =
+              map_flattened_map_args ctxt s (map3 (massage_map_or_map_arg bound_Ts) Us Ts) fs;
+          in
+            Term.list_comb (map', fs')
+          end
+        | NONE => raise AINT_NO_MAP t)
+      | massage_map _ _ _ t = raise AINT_NO_MAP t
+    and massage_map_or_map_arg bound_Ts U T t =
+      if T = U then
+        tap check_no_call t
+      else
+        massage_map bound_Ts U T t
+        handle AINT_NO_MAP _ => massage_mutual_fun bound_Ts U T t;
+
+    fun massage_call bound_Ts U T =
+      massage_let_if_case ctxt has_call (fn bound_Ts => fn t =>
+        if has_call t then
+          (case t of
+            Const (@{const_name prod_case}, _) $ 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'
+            end
+          | t1 $ t2 =>
+            (if has_call t2 then
+              massage_mutual_call bound_Ts U T t
+            else
+              massage_map bound_Ts U T t1 $ t2
+              handle AINT_NO_MAP _ => massage_mutual_call bound_Ts U T t)
+          | Abs (s, T', t') =>
+            Abs (s, T', massage_call (T' :: bound_Ts) (range_type U) (range_type T) t')
+          | _ => massage_mutual_call bound_Ts U T t)
+        else
+          build_map_Inl (T, U) $ t) bound_Ts;
+
+    val T = fastype_of1 (bound_Ts, t);
+  in
+    if has_call t then massage_call bound_Ts U T t else build_map_Inl (T, U) $ t
+  end;
+
+val fold_rev_corec_call = fold_rev_let_if_case;
+
+fun expand_to_ctr_term ctxt s Ts t =
+  (case ctr_sugar_of ctxt s of
+    SOME {ctrs, casex, ...} =>
+    Term.list_comb (mk_case Ts (Type (s, Ts)) casex, map (mk_ctr Ts) ctrs) $ t
+  | NONE => raise Fail "expand_to_ctr_term");
+
+fun expand_corec_code_rhs ctxt has_call bound_Ts t =
+  (case fastype_of1 (bound_Ts, t) of
+    Type (s, Ts) =>
+    massage_let_if_case ctxt has_call (fn _ => fn t =>
+      if can (dest_ctr ctxt s) t then t else expand_to_ctr_term ctxt s Ts t) bound_Ts t
+  | _ => raise Fail "expand_corec_code_rhs");
+
+fun massage_corec_code_rhs ctxt massage_ctr =
+  massage_let_if_case ctxt (K false)
+    (fn bound_Ts => uncurry (massage_ctr bound_Ts) o Term.strip_comb);
+
+fun fold_rev_corec_code_rhs ctxt f =
+  snd ooo 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
+    val (caseT_names, _) = fold_rev_let_if_case ctxt (K (K I)) bound_Ts t ();
+    val ctr_sugars = map (the o ctr_sugar_of ctxt) caseT_names;
+  in
+    (maps #distincts ctr_sugars, maps #discIs ctr_sugars, maps #sel_splits ctr_sugars,
+     maps #sel_split_asms ctr_sugars)
+  end;
+
+fun basic_corec_specs_of ctxt res_T =
+  (case res_T of
+    Type (T_name, _) =>
+    (case Ctr_Sugar.ctr_sugar_of ctxt T_name of
+      NONE => not_codatatype ctxt res_T
+    | SOME {ctrs, discs, selss, ...} =>
+      let
+        val thy = Proof_Context.theory_of ctxt;
+
+        val gfpT = body_type (fastype_of (hd ctrs));
+        val As_rho = tvar_subst thy [gfpT] [res_T];
+        val substA = Term.subst_TVars As_rho;
+
+        fun mk_spec ctr disc sels = {ctr = substA ctr, disc = substA disc, sels = map substA sels};
+      in
+        map3 mk_spec ctrs discs selss
+      end)
+  | _ => not_codatatype ctxt res_T);
+
+fun corec_specs_of bs arg_Ts res_Ts get_indices callssss0 lthy =
+  let
+    val thy = Proof_Context.theory_of lthy;
+
+    val ((missing_res_Ts, perm0_kks,
+          fp_sugars as {nested_bnfs, fp_res = {xtor_co_iterss = dtor_coiters1 :: _, ...},
+            co_inducts = coinduct_thms, ...} :: _, (_, gfp_sugar_thms)), lthy') =
+      nested_to_mutual_fps Greatest_FP bs res_Ts get_indices callssss0 lthy;
+
+    val perm_fp_sugars = sort (int_ord o pairself #index) fp_sugars;
+
+    val indices = map #index fp_sugars;
+    val perm_indices = map #index perm_fp_sugars;
+
+    val perm_ctrss = map (#ctrs o of_fp_sugar #ctr_sugars) perm_fp_sugars;
+    val perm_ctr_Tsss = map (map (binder_types o fastype_of)) perm_ctrss;
+    val perm_gfpTs = map (body_type o fastype_of o hd) perm_ctrss;
+
+    val nn0 = length res_Ts;
+    val nn = length perm_gfpTs;
+    val kks = 0 upto nn - 1;
+    val perm_ns = map length perm_ctr_Tsss;
+
+    val perm_Cs = map (domain_type o body_fun_type o fastype_of o co_rec_of o
+      of_fp_sugar (#xtor_co_iterss o #fp_res)) perm_fp_sugars;
+    val (perm_p_Tss, (perm_q_Tssss, _, perm_f_Tssss, _)) =
+      mk_coiter_fun_arg_types perm_ctr_Tsss perm_Cs perm_ns (co_rec_of dtor_coiters1);
+
+    val (perm_p_hss, h) = indexedd perm_p_Tss 0;
+    val (perm_q_hssss, h') = indexedddd perm_q_Tssss h;
+    val (perm_f_hssss, _) = indexedddd perm_f_Tssss h';
+
+    val fun_arg_hs =
+      flat (map3 flat_corec_preds_predsss_gettersss perm_p_hss perm_q_hssss perm_f_hssss);
+
+    fun unpermute0 perm0_xs = permute_like (op =) perm0_kks kks perm0_xs;
+    fun unpermute perm_xs = permute_like (op =) perm_indices indices perm_xs;
+
+    val coinduct_thmss = map (unpermute0 o conj_dests nn) coinduct_thms;
+
+    val p_iss = map (map (find_index_eq fun_arg_hs)) (unpermute perm_p_hss);
+    val q_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_q_hssss);
+    val f_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_f_hssss);
+
+    val f_Tssss = unpermute perm_f_Tssss;
+    val gfpTs = unpermute perm_gfpTs;
+    val Cs = unpermute perm_Cs;
+
+    val As_rho = tvar_subst thy (take nn0 gfpTs) res_Ts;
+    val Cs_rho = map (fst o dest_TVar) Cs ~~ pad_list HOLogic.unitT nn arg_Ts;
+
+    val substA = Term.subst_TVars As_rho;
+    val substAT = Term.typ_subst_TVars As_rho;
+    val substCT = Term.typ_subst_TVars Cs_rho;
+
+    val perm_Cs' = map substCT perm_Cs;
+
+    fun call_of nullary [] [g_i] [Type (@{type_name fun}, [_, T])] =
+        (if exists_subtype_in Cs T then Nested_Corec
+         else if nullary then Dummy_No_Corec
+         else No_Corec) g_i
+      | call_of _ [q_i] [g_i, g_i'] _ = Mutual_Corec (q_i, g_i, g_i');
+
+    fun mk_ctr_spec ctr disc sels p_ho q_iss f_iss f_Tss discI sel_thms collapse corec_thm
+        disc_corec sel_corecs =
+      let val nullary = not (can dest_funT (fastype_of ctr)) in
+        {ctr = substA ctr, disc = substA disc, sels = map substA sels, pred = p_ho,
+         calls = map3 (call_of nullary) q_iss f_iss f_Tss, discI = discI, sel_thms = sel_thms,
+         collapse = collapse, corec_thm = corec_thm, disc_corec = disc_corec,
+         sel_corecs = sel_corecs}
+      end;
+
+    fun mk_ctr_specs index (ctr_sugars : ctr_sugar list) p_is q_isss f_isss f_Tsss coiter_thmsss
+        disc_coitersss sel_coiterssss =
+      let
+        val ctrs = #ctrs (nth ctr_sugars index);
+        val discs = #discs (nth ctr_sugars index);
+        val selss = #selss (nth ctr_sugars index);
+        val p_ios = map SOME p_is @ [NONE];
+        val discIs = #discIs (nth ctr_sugars index);
+        val sel_thmss = #sel_thmss (nth ctr_sugars index);
+        val collapses = #collapses (nth ctr_sugars index);
+        val corec_thms = co_rec_of (nth coiter_thmsss index);
+        val disc_corecs = co_rec_of (nth disc_coitersss index);
+        val sel_corecss = co_rec_of (nth sel_coiterssss index);
+      in
+        map13 mk_ctr_spec ctrs discs selss p_ios q_isss f_isss f_Tsss discIs sel_thmss collapses
+          corec_thms disc_corecs sel_corecss
+      end;
+
+    fun mk_spec ({T, index, ctr_sugars, co_iterss = coiterss, co_iter_thmsss = coiter_thmsss,
+          disc_co_itersss = disc_coitersss, sel_co_iterssss = sel_coiterssss, ...} : fp_sugar)
+        p_is q_isss f_isss f_Tsss =
+      {corec = mk_co_iter thy Greatest_FP (substAT T) perm_Cs' (co_rec_of (nth coiterss index)),
+       nested_map_idents = map (unfold_thms lthy @{thms id_def} o map_id0_of_bnf) nested_bnfs,
+       nested_map_comps = map map_comp_of_bnf nested_bnfs,
+       ctr_specs = mk_ctr_specs index ctr_sugars p_is q_isss f_isss f_Tsss coiter_thmsss
+         disc_coitersss sel_coiterssss};
+  in
+    ((is_some gfp_sugar_thms, map5 mk_spec fp_sugars p_iss q_issss f_issss f_Tssss, missing_res_Ts,
+      co_induct_of coinduct_thms, strong_co_induct_of coinduct_thms, co_induct_of coinduct_thmss,
+      strong_co_induct_of coinduct_thmss), lthy')
+  end;
+
+val undef_const = Const (@{const_name undefined}, dummyT);
+
+val abs_tuple = HOLogic.tupled_lambda o HOLogic.mk_tuple;
+fun abstract vs =
+  let fun a n (t $ u) = a n t $ a n u
+        | a n (Abs (v, T, b)) = Abs (v, T, a (n + 1) b)
+        | a n t = let val idx = find_index (equal t) vs in
+            if idx < 0 then t else Bound (n + idx) end
+  in a 0 end;
+
+fun mk_prod1 bound_Ts (t, u) =
+  HOLogic.pair_const (fastype_of1 (bound_Ts, t)) (fastype_of1 (bound_Ts, u)) $ t $ u;
+fun mk_tuple1 bound_Ts = the_default HOLogic.unit o try (foldr1 (mk_prod1 bound_Ts));
+
+type coeqn_data_disc = {
+  fun_name: string,
+  fun_T: typ,
+  fun_args: term list,
+  ctr: term,
+  ctr_no: int, (*###*)
+  disc: term,
+  prems: term list,
+  auto_gen: bool,
+  maybe_ctr_rhs: term option,
+  maybe_code_rhs: term option,
+  user_eqn: term
+};
+
+type coeqn_data_sel = {
+  fun_name: string,
+  fun_T: typ,
+  fun_args: term list,
+  ctr: term,
+  sel: term,
+  rhs_term: term,
+  user_eqn: term
+};
+
+datatype coeqn_data =
+  Disc of coeqn_data_disc |
+  Sel of coeqn_data_sel;
+
+fun dissect_coeqn_disc seq fun_names (basic_ctr_specss : basic_corec_ctr_spec list list)
+    maybe_ctr_rhs maybe_code_rhs prems' concl matchedsss =
+  let
+    fun find_subterm p =
+      let (* FIXME \<exists>? *)
+        fun find (t as u $ v) = if p t then SOME t else merge_options (find u, find v)
+          | find t = if p t then SOME t else NONE;
+      in find end;
+
+    val applied_fun = concl
+      |> find_subterm (member ((op =) o apsnd SOME) fun_names o try (fst o dest_Free o head_of))
+      |> the
+      handle Option.Option => primcorec_error_eqn "malformed discriminator formula" concl;
+    val ((fun_name, fun_T), fun_args) = strip_comb applied_fun |>> dest_Free;
+    val SOME basic_ctr_specs = AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name;
+
+    val discs = map #disc basic_ctr_specs;
+    val ctrs = map #ctr basic_ctr_specs;
+    val not_disc = head_of concl = @{term Not};
+    val _ = not_disc andalso length ctrs <> 2 andalso
+      primcorec_error_eqn "negated discriminator for a type with \<noteq> 2 constructors" concl;
+    val disc' = find_subterm (member (op =) discs o head_of) concl;
+    val eq_ctr0 = concl |> perhaps (try HOLogic.dest_not) |> try (HOLogic.dest_eq #> snd)
+        |> (fn SOME t => let val n = find_index (equal t) ctrs in
+          if n >= 0 then SOME n else NONE end | _ => NONE);
+    val _ = is_some disc' orelse is_some eq_ctr0 orelse
+      primcorec_error_eqn "no discriminator in equation" concl;
+    val ctr_no' =
+      if is_none disc' then the eq_ctr0 else find_index (equal (head_of (the disc'))) discs;
+    val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
+    val {ctr, disc, ...} = nth basic_ctr_specs ctr_no;
+
+    val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_;
+    val matchedss = AList.lookup (op =) matchedsss fun_name |> the_default [];
+    val prems = map (abstract (List.rev fun_args)) prems';
+    val real_prems =
+      (if catch_all orelse seq then maps s_not_conj matchedss else []) @
+      (if catch_all then [] else prems);
+
+    val matchedsss' = AList.delete (op =) fun_name matchedsss
+      |> cons (fun_name, if seq then matchedss @ [prems] else matchedss @ [real_prems]);
+
+    val user_eqn =
+      (real_prems, concl)
+      |>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop o abstract (List.rev fun_args)
+      |> curry Logic.list_all (map dest_Free fun_args) o Logic.list_implies;
+  in
+    (Disc {
+      fun_name = fun_name,
+      fun_T = fun_T,
+      fun_args = fun_args,
+      ctr = ctr,
+      ctr_no = ctr_no,
+      disc = disc,
+      prems = real_prems,
+      auto_gen = catch_all,
+      maybe_ctr_rhs = maybe_ctr_rhs,
+      maybe_code_rhs = maybe_code_rhs,
+      user_eqn = user_eqn
+    }, matchedsss')
+  end;
+
+fun dissect_coeqn_sel fun_names (basic_ctr_specss : basic_corec_ctr_spec list list) eqn'
+    maybe_of_spec eqn =
+  let
+    val (lhs, rhs) = HOLogic.dest_eq eqn
+      handle TERM _ =>
+        primcorec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
+    val sel = head_of lhs;
+    val ((fun_name, fun_T), fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst dest_Free
+      handle TERM _ =>
+        primcorec_error_eqn "malformed selector argument in left-hand side" eqn;
+    val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name)
+      handle Option.Option =>
+        primcorec_error_eqn "malformed selector argument in left-hand side" eqn;
+    val {ctr, ...} =
+      (case maybe_of_spec of
+        SOME of_spec => the (find_first (equal of_spec o #ctr) basic_ctr_specs)
+      | NONE => filter (exists (equal sel) o #sels) basic_ctr_specs |> the_single
+          handle List.Empty => primcorec_error_eqn "ambiguous selector - use \"of\"" eqn);
+    val user_eqn = drop_All eqn';
+  in
+    Sel {
+      fun_name = fun_name,
+      fun_T = fun_T,
+      fun_args = fun_args,
+      ctr = ctr,
+      sel = sel,
+      rhs_term = rhs,
+      user_eqn = user_eqn
+    }
+  end;
+
+fun dissect_coeqn_ctr seq fun_names (basic_ctr_specss : basic_corec_ctr_spec list list) eqn'
+    maybe_code_rhs prems concl matchedsss =
+  let
+    val (lhs, rhs) = HOLogic.dest_eq concl;
+    val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
+    val SOME basic_ctr_specs = AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name;
+    val (ctr, ctr_args) = strip_comb (unfold_let rhs);
+    val {disc, sels, ...} = the (find_first (equal ctr o #ctr) basic_ctr_specs)
+      handle Option.Option => primcorec_error_eqn "not a constructor" ctr;
+
+    val disc_concl = betapply (disc, lhs);
+    val (maybe_eqn_data_disc, matchedsss') = if length basic_ctr_specs = 1
+      then (NONE, matchedsss)
+      else apfst SOME (dissect_coeqn_disc seq fun_names basic_ctr_specss
+          (SOME (abstract (List.rev fun_args) rhs)) maybe_code_rhs prems disc_concl matchedsss);
+
+    val sel_concls = sels ~~ ctr_args
+      |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
+
+(*
+val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} concl ^ "\nto\n    \<cdot> " ^
+ (is_some maybe_eqn_data_disc ? K (Syntax.string_of_term @{context} disc_concl ^ "\n    \<cdot> ")) "" ^
+ space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_concls) ^
+ "\nfor premise(s)\n    \<cdot> " ^
+ space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) prems));
+*)
+
+    val eqns_data_sel =
+      map (dissect_coeqn_sel fun_names basic_ctr_specss eqn' (SOME ctr)) sel_concls;
+  in
+    (the_list maybe_eqn_data_disc @ eqns_data_sel, matchedsss')
+  end;
+
+fun dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn' concl matchedsss =
+  let
+    val (lhs, (rhs', rhs)) = HOLogic.dest_eq concl ||> `(expand_corec_code_rhs lthy has_call []);
+    val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
+    val SOME basic_ctr_specs = AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name;
+
+    val cond_ctrs = fold_rev_corec_code_rhs lthy (fn cs => fn ctr => fn _ =>
+        if member ((op =) o apsnd #ctr) basic_ctr_specs ctr
+        then cons (ctr, cs)
+        else primcorec_error_eqn "not a constructor" ctr) [] rhs' []
+      |> AList.group (op =);
+
+    val ctr_premss = (case cond_ctrs of [_] => [[]] | _ => map (s_dnf o snd) cond_ctrs);
+    val ctr_concls = cond_ctrs |> map (fn (ctr, _) =>
+        binder_types (fastype_of ctr)
+        |> map_index (fn (n, T) => massage_corec_code_rhs lthy (fn _ => fn ctr' => fn args =>
+          if ctr' = ctr then nth args n else Const (@{const_name undefined}, T)) [] rhs')
+        |> curry list_comb ctr
+        |> curry HOLogic.mk_eq lhs);
+  in
+    fold_map2 (dissect_coeqn_ctr false fun_names basic_ctr_specss eqn'
+        (SOME (abstract (List.rev fun_args) rhs)))
+      ctr_premss ctr_concls matchedsss
+  end;
+
+fun dissect_coeqn lthy seq has_call fun_names (basic_ctr_specss : basic_corec_ctr_spec list list)
+    eqn' maybe_of_spec matchedsss =
+  let
+    val eqn = drop_All eqn'
+      handle TERM _ => primcorec_error_eqn "malformed function equation" eqn';
+    val (prems, concl) = Logic.strip_horn eqn
+      |> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop;
+
+    val head = concl
+      |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
+      |> head_of;
+
+    val maybe_rhs = concl |> perhaps (try HOLogic.dest_not) |> try (snd o HOLogic.dest_eq);
+
+    val discs = maps (map #disc) basic_ctr_specss;
+    val sels = maps (maps #sels) basic_ctr_specss;
+    val ctrs = maps (map #ctr) basic_ctr_specss;
+  in
+    if member (op =) discs head orelse
+      is_some maybe_rhs andalso
+        member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then
+      dissect_coeqn_disc seq fun_names basic_ctr_specss NONE NONE prems concl matchedsss
+      |>> single
+    else if member (op =) sels head then
+      ([dissect_coeqn_sel fun_names basic_ctr_specss eqn' maybe_of_spec concl], matchedsss)
+    else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) andalso
+      member (op =) ctrs (head_of (unfold_let (the maybe_rhs))) then
+      dissect_coeqn_ctr seq fun_names basic_ctr_specss eqn' NONE prems concl matchedsss
+    else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) andalso
+      null prems then
+      dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn' concl matchedsss
+      |>> flat
+    else
+      primcorec_error_eqn "malformed function equation" eqn
+  end;
+
+fun build_corec_arg_disc (ctr_specs : corec_ctr_spec list)
+    ({fun_args, ctr_no, prems, ...} : coeqn_data_disc) =
+  if is_none (#pred (nth ctr_specs ctr_no)) then I else
+    s_conjs prems
+    |> curry subst_bounds (List.rev fun_args)
+    |> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args)
+    |> K |> nth_map (the (#pred (nth ctr_specs ctr_no)));
+
+fun build_corec_arg_no_call (sel_eqns : coeqn_data_sel list) sel =
+  find_first (equal sel o #sel) sel_eqns
+  |> try (fn SOME {fun_args, rhs_term, ...} => abs_tuple fun_args rhs_term)
+  |> the_default undef_const
+  |> K;
+
+fun build_corec_args_mutual_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
+  (case find_first (equal sel o #sel) sel_eqns of
+    NONE => (I, I, I)
+  | SOME {fun_args, rhs_term, ... } =>
+    let
+      val bound_Ts = List.rev (map fastype_of fun_args);
+      fun rewrite_stop _ t = if has_call t then @{term False} else @{term True};
+      fun rewrite_end _ t = if has_call t then undef_const else t;
+      fun rewrite_cont bound_Ts t =
+        if has_call t then mk_tuple1 bound_Ts (snd (strip_comb t)) else undef_const;
+      fun massage f _ = massage_mutual_corec_call lthy has_call f bound_Ts rhs_term
+        |> abs_tuple fun_args;
+    in
+      (massage rewrite_stop, massage rewrite_end, massage rewrite_cont)
+    end);
+
+fun build_corec_arg_nested_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
+  (case find_first (equal sel o #sel) sel_eqns of
+    NONE => I
+  | SOME {fun_args, rhs_term, ...} =>
+    let
+      val bound_Ts = List.rev (map fastype_of fun_args);
+      fun rewrite bound_Ts U T (Abs (v, V, b)) = Abs (v, V, rewrite (V :: bound_Ts) U T b)
+        | rewrite bound_Ts U T (t as _ $ _) =
+          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
+              map (rewrite bound_Ts U T) vs |> chop 1 |>> HOLogic.mk_split o the_single |> list_comb
+            else
+              list_comb (rewrite bound_Ts U T u, map (rewrite bound_Ts U T) vs)
+          end
+        | rewrite _ U T t =
+          if is_Free t andalso has_call t then Inr_const U T $ HOLogic.unit else t;
+      fun massage t =
+        rhs_term
+        |> massage_nested_corec_call lthy has_call rewrite bound_Ts (range_type (fastype_of t))
+        |> abs_tuple fun_args;
+    in
+      massage
+    end);
+
+fun build_corec_args_sel lthy has_call (all_sel_eqns : coeqn_data_sel list)
+    (ctr_spec : corec_ctr_spec) =
+  (case filter (equal (#ctr ctr_spec) o #ctr) all_sel_eqns of
+    [] => I
+  | sel_eqns =>
+    let
+      val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec;
+      val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list;
+      val mutual_calls' = map_filter (try (apsnd (fn Mutual_Corec n => n))) sel_call_list;
+      val nested_calls' = map_filter (try (apsnd (fn Nested_Corec n => n))) sel_call_list;
+    in
+      I
+      #> fold (fn (sel, n) => nth_map n (build_corec_arg_no_call sel_eqns sel)) no_calls'
+      #> fold (fn (sel, (q, g, h)) =>
+        let val (fq, fg, fh) = build_corec_args_mutual_call lthy has_call sel_eqns sel in
+          nth_map q fq o nth_map g fg o nth_map h fh end) mutual_calls'
+      #> fold (fn (sel, n) => nth_map n
+        (build_corec_arg_nested_call lthy has_call sel_eqns sel)) nested_calls'
+    end);
+
+fun build_codefs lthy bs mxs has_call arg_Tss (corec_specs : corec_spec list)
+    (disc_eqnss : coeqn_data_disc list list) (sel_eqnss : coeqn_data_sel list list) =
+  let
+    val corecs = map #corec corec_specs;
+    val ctr_specss = map #ctr_specs corec_specs;
+    val corec_args = hd corecs
+      |> fst o split_last o binder_types o fastype_of
+      |> map (Const o pair @{const_name undefined})
+      |> fold2 (fold o build_corec_arg_disc) ctr_specss disc_eqnss
+      |> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss;
+    fun currys [] t = t
+      | currys Ts t = t $ mk_tuple1 (List.rev Ts) (map Bound (length Ts - 1 downto 0))
+          |> fold_rev (Term.abs o pair Name.uu) Ts;
+
+(*
+val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
+ space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
+*)
+
+    val exclss' =
+      disc_eqnss
+      |> map (map (fn x => (#fun_args x, #ctr_no x, #prems x, #auto_gen x))
+        #> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs [])
+        #> maps (uncurry (map o pair)
+          #> map (fn ((fun_args, c, x, a), (_, c', y, a')) =>
+              ((c, c', a orelse a'), (x, s_not (s_conjs y)))
+            ||> apfst (map HOLogic.mk_Trueprop) o apsnd HOLogic.mk_Trueprop
+            ||> Logic.list_implies
+            ||> curry Logic.list_all (map dest_Free fun_args))))
+  in
+    map (list_comb o rpair corec_args) corecs
+    |> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss
+    |> map2 currys arg_Tss
+    |> Syntax.check_terms lthy
+    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
+      bs mxs
+    |> rpair exclss'
+  end;
+
+fun mk_real_disc_eqns fun_binding arg_Ts ({ctr_specs, ...} : corec_spec)
+    (sel_eqns : coeqn_data_sel list) (disc_eqns : coeqn_data_disc list) =
+  if length disc_eqns <> length ctr_specs - 1 then disc_eqns else
+    let
+      val n = 0 upto length ctr_specs
+        |> the o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns));
+      val fun_args = (try (#fun_args o hd) disc_eqns, try (#fun_args o hd) sel_eqns)
+        |> the_default (map (curry Free Name.uu) arg_Ts) o merge_options;
+      val extra_disc_eqn = {
+        fun_name = Binding.name_of fun_binding,
+        fun_T = arg_Ts ---> body_type (fastype_of (#ctr (hd ctr_specs))),
+        fun_args = fun_args,
+        ctr = #ctr (nth ctr_specs n),
+        ctr_no = n,
+        disc = #disc (nth ctr_specs n),
+        prems = maps (s_not_conj o #prems) disc_eqns,
+        auto_gen = true,
+        maybe_ctr_rhs = NONE,
+        maybe_code_rhs = NONE,
+        user_eqn = undef_const};
+    in
+      chop n disc_eqns ||> cons extra_disc_eqn |> (op @)
+    end;
+
+fun find_corec_calls ctxt has_call basic_ctr_specs ({ctr, sel, rhs_term, ...} : coeqn_data_sel) =
+  let
+    val sel_no = find_first (equal ctr o #ctr) basic_ctr_specs
+      |> find_index (equal sel) o #sels o the;
+    fun find t = if has_call t then snd (fold_rev_corec_call ctxt (K cons) [] t []) else [];
+  in
+    find rhs_term
+    |> K |> nth_map sel_no |> AList.map_entry (op =) ctr
+  end;
+
+fun add_primcorec_ursive maybe_tac seq fixes specs maybe_of_specs lthy =
+  let
+    val thy = Proof_Context.theory_of lthy;
+
+    val (bs, mxs) = map_split (apfst fst) fixes;
+    val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list;
+
+    val _ = (case filter_out (fn (_, T) => Sign.of_sort thy (T, HOLogic.typeS)) (bs ~~ arg_Ts) of
+        [] => ()
+      | (b, _) :: _ => primcorec_error ("type of " ^ Binding.print b ^ " contains top sort"));
+
+    val fun_names = map Binding.name_of bs;
+    val basic_ctr_specss = map (basic_corec_specs_of lthy) res_Ts;
+    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
+    val eqns_data =
+      fold_map2 (dissect_coeqn lthy seq has_call fun_names basic_ctr_specss) (map snd specs)
+        maybe_of_specs []
+      |> flat o fst;
+
+    val callssss =
+      map_filter (try (fn Sel x => x)) eqns_data
+      |> partition_eq ((op =) o pairself #fun_name)
+      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
+      |> map (flat o snd)
+      |> map2 (fold o find_corec_calls lthy has_call) basic_ctr_specss
+      |> map2 (curry (op |>)) (map (map (fn {ctr, sels, ...} =>
+        (ctr, map (K []) sels))) basic_ctr_specss);
+
+(*
+val _ = tracing ("callssss = " ^ @{make_string} callssss);
+*)
+
+    val ((n2m, corec_specs', _, coinduct_thm, strong_coinduct_thm, coinduct_thms,
+          strong_coinduct_thms), lthy') =
+      corec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
+    val actual_nn = length bs;
+    val corec_specs = take actual_nn corec_specs'; (*###*)
+    val ctr_specss = map #ctr_specs corec_specs;
+
+    val disc_eqnss' = map_filter (try (fn Disc x => x)) eqns_data
+      |> partition_eq ((op =) o pairself #fun_name)
+      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
+      |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
+    val _ = disc_eqnss' |> map (fn x =>
+      let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse
+        primcorec_error_eqns "excess discriminator formula in definition"
+          (maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
+
+    val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data
+      |> partition_eq ((op =) o pairself #fun_name)
+      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
+      |> map (flat o snd);
+
+    val arg_Tss = map (binder_types o snd o fst) fixes;
+    val disc_eqnss = map5 mk_real_disc_eqns bs arg_Tss corec_specs sel_eqnss disc_eqnss';
+    val (defs, exclss') =
+      build_codefs lthy' bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss;
+
+    fun excl_tac (c, c', a) =
+      if a orelse c = c' orelse seq then SOME (K (HEADGOAL (mk_primcorec_assumption_tac lthy [])))
+      else maybe_tac;
+
+(*
+val _ = tracing ("exclusiveness properties:\n    \<cdot> " ^
+ space_implode "\n    \<cdot> " (maps (map (Syntax.string_of_term lthy o snd)) exclss'));
+*)
+
+    val exclss'' = exclss' |> map (map (fn (idx, t) =>
+      (idx, (Option.map (Goal.prove lthy [] [] t #> Thm.close_derivation) (excl_tac idx), t))));
+    val taut_thmss = map (map (apsnd (the o fst)) o filter (is_some o fst o snd)) exclss'';
+    val (goal_idxss, goalss) = exclss''
+      |> map (map (apsnd (rpair [] o snd)) o filter (is_none o fst o snd))
+      |> split_list o map split_list;
+
+    fun prove thmss' def_thms' lthy =
+      let
+        val def_thms = map (snd o snd) def_thms';
+
+        val exclss' = map (op ~~) (goal_idxss ~~ thmss');
+        fun mk_exclsss excls n =
+          (excls, map (fn k => replicate k [TrueI] @ replicate (n - k) []) (0 upto n - 1))
+          |-> fold (fn ((c, c', _), thm) => nth_map c (nth_map c' (K [thm])));
+        val exclssss = (exclss' ~~ taut_thmss |> map (op @), fun_names ~~ corec_specs)
+          |-> map2 (fn excls => fn (_, {ctr_specs, ...}) => mk_exclsss excls (length ctr_specs));
+
+        fun prove_disc ({ctr_specs, ...} : corec_spec) exclsss
+            ({fun_name, fun_T, fun_args, ctr_no, prems, ...} : coeqn_data_disc) =
+          if Term.aconv_untyped (#disc (nth ctr_specs ctr_no), @{term "\<lambda>x. x = x"}) then
+            []
+          else
+            let
+              val {disc_corec, ...} = nth ctr_specs ctr_no;
+              val k = 1 + ctr_no;
+              val m = length prems;
+              val t =
+                list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
+                |> curry betapply (#disc (nth ctr_specs ctr_no)) (*###*)
+                |> HOLogic.mk_Trueprop
+                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
+                |> curry Logic.list_all (map dest_Free fun_args);
+            in
+              if prems = [@{term False}] then [] else
+              mk_primcorec_disc_tac lthy def_thms disc_corec k m exclsss
+              |> K |> Goal.prove lthy [] [] t
+              |> Thm.close_derivation
+              |> pair (#disc (nth ctr_specs ctr_no))
+              |> single
+            end;
+
+        fun prove_sel ({nested_map_idents, nested_map_comps, ctr_specs, ...} : corec_spec)
+            (disc_eqns : coeqn_data_disc list) exclsss
+            ({fun_name, fun_T, fun_args, ctr, sel, rhs_term, ...} : coeqn_data_sel) =
+          let
+            val SOME ctr_spec = find_first (equal ctr o #ctr) ctr_specs;
+            val ctr_no = find_index (equal ctr o #ctr) ctr_specs;
+            val prems = the_default (maps (s_not_conj o #prems) disc_eqns)
+                (find_first (equal ctr_no o #ctr_no) disc_eqns |> Option.map #prems);
+            val sel_corec = find_index (equal sel) (#sels ctr_spec)
+              |> nth (#sel_corecs ctr_spec);
+            val k = 1 + ctr_no;
+            val m = length prems;
+            val t =
+              list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
+              |> curry betapply sel
+              |> rpair (abstract (List.rev fun_args) rhs_term)
+              |> 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;
+          in
+            mk_primcorec_sel_tac lthy def_thms distincts sel_splits sel_split_asms nested_map_idents
+              nested_map_comps sel_corec k m exclsss
+            |> K |> Goal.prove lthy [] [] t
+            |> Thm.close_derivation
+            |> pair sel
+          end;
+
+        fun prove_ctr disc_alist sel_alist (disc_eqns : coeqn_data_disc list)
+            (sel_eqns : coeqn_data_sel list) ({ctr, disc, sels, collapse, ...} : corec_ctr_spec) =
+          (* don't try to prove theorems when some sel_eqns are missing *)
+          if not (exists (equal ctr o #ctr) disc_eqns)
+              andalso not (exists (equal ctr o #ctr) sel_eqns)
+            orelse
+              filter (equal ctr o #ctr) sel_eqns
+              |> fst o finds ((op =) o apsnd #sel) sels
+              |> exists (null o snd)
+          then [] else
+            let
+              val (fun_name, fun_T, fun_args, prems, maybe_rhs) =
+                (find_first (equal ctr o #ctr) disc_eqns, find_first (equal ctr o #ctr) sel_eqns)
+                |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #prems x,
+                  #maybe_ctr_rhs x))
+                ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, [], NONE))
+                |> the o merge_options;
+              val m = length prems;
+              val t = (if is_some maybe_rhs then the maybe_rhs else
+                  filter (equal ctr o #ctr) sel_eqns
+                  |> fst o finds ((op =) o apsnd #sel) sels
+                  |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x)) #-> abstract)
+                  |> curry list_comb ctr)
+                |> curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
+                  map Bound (length fun_args - 1 downto 0)))
+                |> HOLogic.mk_Trueprop
+                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
+                |> curry Logic.list_all (map dest_Free fun_args);
+              val maybe_disc_thm = AList.lookup (op =) disc_alist disc;
+              val sel_thms = map snd (filter (member (op =) sels o fst) sel_alist);
+            in
+              if prems = [@{term False}] then [] else
+                mk_primcorec_ctr_of_dtr_tac lthy m collapse maybe_disc_thm sel_thms
+                |> K |> Goal.prove lthy [] [] t
+                |> Thm.close_derivation
+                |> pair ctr
+                |> single
+            end;
+
+        fun prove_code disc_eqns sel_eqns ctr_alist ctr_specs =
+          let
+            val (fun_name, fun_T, fun_args, maybe_rhs) =
+              (find_first (member (op =) (map #ctr ctr_specs) o #ctr) disc_eqns,
+               find_first (member (op =) (map #ctr ctr_specs) o #ctr) sel_eqns)
+              |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #maybe_code_rhs x))
+              ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, NONE))
+              |> the o merge_options;
+
+            val bound_Ts = List.rev (map fastype_of fun_args);
+
+            val lhs = list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0));
+            val maybe_rhs_info =
+              (case maybe_rhs of
+                SOME rhs =>
+                let
+                  val raw_rhs = expand_corec_code_rhs lthy has_call bound_Ts rhs;
+                  val cond_ctrs =
+                    fold_rev_corec_code_rhs lthy (K oo (cons oo pair)) bound_Ts raw_rhs [];
+                  val ctr_thms = map (the o AList.lookup (op =) ctr_alist o snd) cond_ctrs;
+                in SOME (rhs, raw_rhs, ctr_thms) end
+              | NONE =>
+                let
+                  fun prove_code_ctr {ctr, sels, ...} =
+                    if not (exists (equal ctr o fst) ctr_alist) then NONE else
+                      let
+                        val prems = find_first (equal ctr o #ctr) disc_eqns
+                          |> Option.map #prems |> the_default [];
+                        val t =
+                          filter (equal ctr o #ctr) sel_eqns
+                          |> fst o finds ((op =) o apsnd #sel) sels
+                          |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x))
+                            #-> abstract)
+                          |> curry list_comb ctr;
+                      in
+                        SOME (prems, t)
+                      end;
+                  val maybe_ctr_conds_argss = map prove_code_ctr ctr_specs;
+                in
+                  if exists is_none maybe_ctr_conds_argss then NONE else
+                    let
+                      val rhs = fold_rev (fn SOME (prems, u) => fn t => mk_If (s_conjs prems) u t)
+                        maybe_ctr_conds_argss
+                        (Const (@{const_name Code.abort}, @{typ String.literal} -->
+                            (@{typ unit} --> body_type fun_T) --> body_type fun_T) $
+                          HOLogic.mk_literal fun_name $
+                          absdummy @{typ unit} (incr_boundvars 1 lhs));
+                    in SOME (rhs, rhs, map snd ctr_alist) end
+                end);
+          in
+            (case maybe_rhs_info of
+              NONE => []
+            | SOME (rhs, raw_rhs, ctr_thms) =>
+              let
+                val ms = map (Logic.count_prems o prop_of) ctr_thms;
+                val (raw_t, t) = (raw_rhs, rhs)
+                  |> pairself
+                    (curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
+                      map Bound (length fun_args - 1 downto 0)))
+                    #> HOLogic.mk_Trueprop
+                    #> 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;
+
+                val raw_code_thm = mk_primcorec_raw_code_of_ctr_tac lthy distincts discIs sel_splits
+                    sel_split_asms ms ctr_thms
+                  |> K |> Goal.prove lthy [] [] raw_t
+                  |> Thm.close_derivation;
+              in
+                mk_primcorec_code_of_raw_code_tac lthy distincts sel_splits raw_code_thm
+                |> K |> Goal.prove lthy [] [] t
+                |> Thm.close_derivation
+                |> single
+              end)
+          end;
+
+        val disc_alists = map3 (maps oo prove_disc) corec_specs exclssss disc_eqnss;
+        val sel_alists = map4 (map ooo prove_sel) corec_specs disc_eqnss exclssss sel_eqnss;
+        val disc_thmss = map (map snd) disc_alists;
+        val sel_thmss = map (map snd) sel_alists;
+
+        val ctr_alists = map5 (maps oooo prove_ctr) disc_alists sel_alists disc_eqnss sel_eqnss
+          ctr_specss;
+        val ctr_thmss = map (map snd) ctr_alists;
+
+        val code_thmss = map4 prove_code disc_eqnss sel_eqnss ctr_alists ctr_specss;
+
+        val simp_thmss = map2 append disc_thmss sel_thmss
+
+        val common_name = mk_common_name fun_names;
+
+        val notes =
+          [(coinductN, map (if n2m then single else K []) coinduct_thms, []),
+           (codeN, code_thmss, code_nitpicksimp_attrs),
+           (ctrN, ctr_thmss, []),
+           (discN, disc_thmss, simp_attrs),
+           (selN, sel_thmss, simp_attrs),
+           (simpsN, simp_thmss, []),
+           (strong_coinductN, map (if n2m then single else K []) strong_coinduct_thms, [])]
+          |> maps (fn (thmN, thmss, attrs) =>
+            map2 (fn fun_name => fn thms =>
+                ((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])]))
+              fun_names (take actual_nn thmss))
+          |> filter_out (null o fst o hd o snd);
+
+        val common_notes =
+          [(coinductN, if n2m then [coinduct_thm] else [], []),
+           (strong_coinductN, if n2m then [strong_coinduct_thm] else [], [])]
+          |> filter_out (null o #2)
+          |> map (fn (thmN, thms, attrs) =>
+            ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
+      in
+        lthy |> Local_Theory.notes (notes @ common_notes) |> snd
+      end;
+
+    fun after_qed thmss' = fold_map Local_Theory.define defs #-> prove thmss';
+  in
+    (goalss, after_qed, lthy')
+  end;
+
+fun add_primcorec_ursive_cmd maybe_tac seq (raw_fixes, raw_specs') lthy =
+  let
+    val (raw_specs, maybe_of_specs) =
+      split_list raw_specs' ||> map (Option.map (Syntax.read_term lthy));
+    val ((fixes, specs), _) = Specification.read_spec raw_fixes raw_specs lthy;
+  in
+    add_primcorec_ursive maybe_tac seq fixes specs maybe_of_specs lthy
+    handle ERROR str => primcorec_error str
+  end
+  handle Primcorec_Error (str, eqns) =>
+    if null eqns
+    then error ("primcorec error:\n  " ^ str)
+    else error ("primcorec error:\n  " ^ str ^ "\nin\n  " ^
+      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
+
+val add_primcorecursive_cmd = (fn (goalss, after_qed, lthy) =>
+  lthy
+  |> Proof.theorem NONE after_qed goalss
+  |> Proof.refine (Method.primitive_text I)
+  |> Seq.hd) ooo add_primcorec_ursive_cmd NONE;
+
+val add_primcorec_cmd = (fn (goalss, after_qed, lthy) =>
+  lthy
+  |> after_qed (map (fn [] => []
+      | _ => primcorec_error "need exclusiveness proofs - use primcorecursive instead of primcorec")
+    goalss)) ooo add_primcorec_ursive_cmd (SOME (fn {context = ctxt, ...} => auto_tac ctxt));
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_gfp_rec_sugar_tactics.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -0,0 +1,135 @@
+(*  Title:      HOL/BNF/Tools/bnf_gfp_rec_sugar_tactics.ML
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2013
+
+Tactics for corecursor sugar.
+*)
+
+signature BNF_GFP_REC_SUGAR_TACTICS =
+sig
+  val mk_primcorec_assumption_tac: Proof.context -> thm list -> int -> tactic
+  val mk_primcorec_code_of_raw_code_tac: Proof.context -> thm list -> thm list -> thm -> tactic
+  val mk_primcorec_ctr_of_dtr_tac: Proof.context -> int -> thm -> thm option -> thm list -> tactic
+  val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list ->
+    tactic
+  val mk_primcorec_raw_code_of_ctr_tac: Proof.context -> thm list -> thm list -> thm list ->
+    thm list -> int list -> thm list -> tactic
+  val mk_primcorec_sel_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
+    thm list -> thm list -> thm -> int -> int -> thm list list list -> tactic
+end;
+
+structure BNF_GFP_Rec_Sugar_Tactics : BNF_GFP_REC_SUGAR_TACTICS =
+struct
+
+open BNF_Util
+open BNF_Tactics
+
+val falseEs = @{thms not_TrueE FalseE};
+val Let_def = @{thm Let_def};
+val neq_eq_eq_contradict = @{thm neq_eq_eq_contradict};
+val split_if = @{thm split_if};
+val split_if_asm = @{thm split_if_asm};
+val split_connectI = @{thms allI impI conjI};
+
+fun mk_primcorec_assumption_tac ctxt discIs =
+  SELECT_GOAL (unfold_thms_tac ctxt
+      @{thms not_not not_False_eq_True not_True_eq_False de_Morgan_conj de_Morgan_disj} THEN
+    SOLVE (HEADGOAL (REPEAT o (rtac refl ORELSE' atac ORELSE' etac conjE ORELSE'
+    eresolve_tac falseEs ORELSE'
+    resolve_tac @{thms TrueI conjI disjI1 disjI2} ORELSE'
+    dresolve_tac discIs THEN' atac ORELSE'
+    etac notE THEN' atac ORELSE'
+    etac disjE))));
+
+fun mk_primcorec_same_case_tac m =
+  HEADGOAL (if m = 0 then rtac TrueI
+    else REPEAT_DETERM_N (m - 1) o (rtac conjI THEN' atac) THEN' atac);
+
+fun mk_primcorec_different_case_tac ctxt m excl =
+  HEADGOAL (if m = 0 then mk_primcorec_assumption_tac ctxt []
+    else dtac excl THEN' (REPEAT_DETERM_N (m - 1) o atac) THEN' mk_primcorec_assumption_tac ctxt []);
+
+fun mk_primcorec_cases_tac ctxt k m exclsss =
+  let val n = length exclsss in
+    EVERY (map (fn [] => if k = n then all_tac else mk_primcorec_same_case_tac m
+        | [excl] => mk_primcorec_different_case_tac ctxt m excl)
+      (take k (nth exclsss (k - 1))))
+  end;
+
+fun mk_primcorec_prelude ctxt defs thm =
+  unfold_thms_tac ctxt defs THEN HEADGOAL (rtac thm) THEN
+  unfold_thms_tac ctxt @{thms Let_def split};
+
+fun mk_primcorec_disc_tac ctxt defs disc_corec k m exclsss =
+  mk_primcorec_prelude ctxt defs disc_corec THEN mk_primcorec_cases_tac ctxt k m exclsss;
+
+fun mk_primcorec_sel_tac ctxt defs distincts splits split_asms map_idents map_comps f_sel k m
+    exclsss =
+  mk_primcorec_prelude ctxt defs (f_sel RS trans) THEN
+  mk_primcorec_cases_tac ctxt k m exclsss THEN
+  HEADGOAL (REPEAT_DETERM o (rtac refl ORELSE' rtac ext ORELSE'
+    eresolve_tac falseEs ORELSE'
+    resolve_tac split_connectI ORELSE'
+    Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
+    Splitter.split_tac (split_if :: splits) ORELSE'
+    eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
+    etac notE THEN' atac ORELSE'
+    (CHANGED o SELECT_GOAL (unfold_thms_tac ctxt
+      (@{thms id_def o_def split_def sum.cases} @ map_comps @ map_idents)))));
+
+fun mk_primcorec_ctr_of_dtr_tac ctxt m collapse maybe_disc_f sel_fs =
+  HEADGOAL (rtac ((if null sel_fs then collapse else collapse RS sym) RS trans) THEN'
+    (the_default (K all_tac) (Option.map rtac maybe_disc_f)) THEN' REPEAT_DETERM_N m o atac) THEN
+  unfold_thms_tac ctxt (Let_def :: sel_fs) THEN HEADGOAL (rtac refl);
+
+fun inst_split_eq ctxt split =
+  (case prop_of split of
+    @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ (Var (_, Type (_, [T, _])) $ _) $ _) =>
+    let
+      val s = Name.uu;
+      val eq = Abs (Name.uu, T, HOLogic.mk_eq (Free (s, T), Bound 0));
+      val split' = Drule.instantiate' [] [SOME (certify ctxt eq)] split;
+    in
+      Thm.generalize ([], [s]) (Thm.maxidx_of split' + 1) split'
+    end
+  | _ => split);
+
+fun distinct_in_prems_tac distincts =
+  eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac;
+
+(* TODO: reduce code duplication with selector tactic above *)
+fun mk_primcorec_raw_code_of_ctr_single_tac ctxt distincts discIs splits split_asms m f_ctr =
+  let
+    val splits' =
+      map (fn th => th RS iffD2) (@{thm split_if_eq2} :: map (inst_split_eq ctxt) splits)
+  in
+    HEADGOAL (REPEAT o (resolve_tac (splits' @ split_connectI))) THEN
+    mk_primcorec_prelude ctxt [] (f_ctr RS trans) THEN
+    HEADGOAL ((REPEAT_DETERM_N m o mk_primcorec_assumption_tac ctxt discIs) THEN'
+      SELECT_GOAL (SOLVE (HEADGOAL (REPEAT_DETERM o
+      (rtac refl ORELSE' atac ORELSE'
+       resolve_tac (@{thm Code.abort_def} :: split_connectI) ORELSE'
+       Splitter.split_tac (split_if :: splits) ORELSE'
+       Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
+       mk_primcorec_assumption_tac ctxt discIs ORELSE'
+       distinct_in_prems_tac distincts ORELSE'
+       (TRY o dresolve_tac discIs) THEN' etac notE THEN' atac)))))
+  end;
+
+fun mk_primcorec_raw_code_of_ctr_tac ctxt distincts discIs splits split_asms ms f_ctrs =
+  EVERY (map2 (mk_primcorec_raw_code_of_ctr_single_tac ctxt distincts discIs splits split_asms) ms
+    f_ctrs) THEN
+  IF_UNSOLVED (unfold_thms_tac ctxt @{thms Code.abort_def} THEN
+    HEADGOAL (REPEAT_DETERM o resolve_tac (refl :: split_connectI)));
+
+fun mk_primcorec_code_of_raw_code_tac ctxt distincts splits raw =
+  HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN'
+    SELECT_GOAL (unfold_thms_tac ctxt [Let_def]) THEN' REPEAT_DETERM o
+    (rtac refl ORELSE' atac ORELSE'
+     resolve_tac split_connectI ORELSE'
+     Splitter.split_tac (split_if :: splits) ORELSE'
+     distinct_in_prems_tac distincts ORELSE'
+     rtac sym THEN' atac ORELSE'
+     etac notE THEN' atac));
+
+end;
--- a/src/HOL/BNF/Tools/bnf_lfp.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Tools/bnf_lfp.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -22,7 +22,7 @@
 open BNF_Comp
 open BNF_FP_Util
 open BNF_FP_Def_Sugar
-open BNF_FP_Rec_Sugar
+open BNF_LFP_Rec_Sugar
 open BNF_LFP_Util
 open BNF_LFP_Tactics
 
--- a/src/HOL/BNF/Tools/bnf_lfp_compat.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Tools/bnf_lfp_compat.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -57,8 +57,10 @@
     val fpTs as fpT1 :: _ = map (fn s => Type (s, As)) fpT_names';
 
     fun add_nested_types_of (T as Type (s, _)) seen =
-      if member (op =) seen T orelse s = @{type_name fun} then
+      if member (op =) seen T then
         seen
+      else if s = @{type_name fun} then
+        (warning "Partial support for recursion through functions -- 'primrec' will fail"; seen)
       else
         (case try lfp_sugar_of s of
           SOME ({T = T0, fp_res = {Ts = mutual_Ts0, ...}, ctr_sugars, ...}) =>
@@ -91,11 +93,13 @@
     val nn = length Ts;
     val get_indices = K [];
     val fp_sugars0 = if nn = 1 then [fp_sugar0] else map (lfp_sugar_of o fst o dest_Type) Ts;
-    val callssss = pad_and_indexify_calls fp_sugars0 nn [];
-    val has_nested = nn > nn_fp;
+    val callssss = map (fn fp_sugar0 => indexify_callsss fp_sugar0 []) fp_sugars0;
 
     val ((fp_sugars, (lfp_sugar_thms, _)), lthy) =
-      mutualize_fp_sugars has_nested Least_FP compat_bs Ts get_indices callssss fp_sugars0 lthy;
+      if nn > nn_fp then
+        mutualize_fp_sugars Least_FP compat_bs Ts get_indices callssss fp_sugars0 lthy
+      else
+        ((fp_sugars0, (NONE, NONE)), lthy);
 
     val {ctr_sugars, co_inducts = [induct], co_iterss, co_iter_thmsss = iter_thmsss, ...} :: _ =
       fp_sugars;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_lfp_rec_sugar.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -0,0 +1,604 @@
+(*  Title:      HOL/BNF/Tools/bnf_lfp_rec_sugar.ML
+    Author:     Lorenz Panny, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2013
+
+Recursor sugar.
+*)
+
+signature BNF_LFP_REC_SUGAR =
+sig
+  val add_primrec: (binding * typ option * mixfix) list ->
+    (Attrib.binding * term) list -> local_theory -> (term list * thm list list) * local_theory
+  val add_primrec_cmd: (binding * string option * mixfix) list ->
+    (Attrib.binding * string) list -> local_theory -> (term list * thm list list) * local_theory
+  val add_primrec_global: (binding * typ option * mixfix) list ->
+    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
+  val add_primrec_overloaded: (string * (string * typ) * bool) list ->
+    (binding * typ option * mixfix) list ->
+    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
+  val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
+    local_theory -> (string list * (term list * (int list list * thm list list))) * local_theory
+end;
+
+structure BNF_LFP_Rec_Sugar : BNF_LFP_REC_SUGAR =
+struct
+
+open Ctr_Sugar
+open BNF_Util
+open BNF_Tactics
+open BNF_Def
+open BNF_FP_Util
+open BNF_FP_Def_Sugar
+open BNF_FP_N2M_Sugar
+open BNF_FP_Rec_Sugar_Util
+
+val nitpicksimp_attrs = @{attributes [nitpick_simp]};
+val simp_attrs = @{attributes [simp]};
+val code_nitpicksimp_simp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs @ simp_attrs;
+
+exception Primrec_Error of string * term list;
+
+fun primrec_error str = raise Primrec_Error (str, []);
+fun primrec_error_eqn str eqn = raise Primrec_Error (str, [eqn]);
+fun primrec_error_eqns str eqns = raise Primrec_Error (str, eqns);
+
+datatype rec_call =
+  No_Rec of int * typ |
+  Mutual_Rec of (int * typ) * (int * typ) |
+  Nested_Rec of int * typ;
+
+type rec_ctr_spec =
+  {ctr: term,
+   offset: int,
+   calls: rec_call list,
+   rec_thm: thm};
+
+type rec_spec =
+  {recx: term,
+   nested_map_idents: thm list,
+   nested_map_comps: thm list,
+   ctr_specs: rec_ctr_spec list};
+
+exception AINT_NO_MAP of term;
+
+fun ill_formed_rec_call ctxt t =
+  error ("Ill-formed recursive call: " ^ quote (Syntax.string_of_term ctxt t));
+fun invalid_map ctxt t =
+  error ("Invalid map function in " ^ quote (Syntax.string_of_term ctxt t));
+fun unexpected_rec_call ctxt t =
+  error ("Unexpected recursive call: " ^ quote (Syntax.string_of_term ctxt t));
+
+fun massage_nested_rec_call ctxt has_call raw_massage_fun bound_Ts y y' =
+  let
+    fun check_no_call t = if has_call t then unexpected_rec_call ctxt t else ();
+
+    val typof = curry fastype_of1 bound_Ts;
+    val build_map_fst = build_map ctxt (fst_const o fst);
+
+    val yT = typof y;
+    val yU = typof y';
+
+    fun y_of_y' () = build_map_fst (yU, yT) $ y';
+    val elim_y = Term.map_aterms (fn t => if t = y then y_of_y' () else t);
+
+    fun massage_mutual_fun U T t =
+      (case t of
+        Const (@{const_name comp}, _) $ t1 $ t2 =>
+        mk_comp bound_Ts (tap check_no_call t1, massage_mutual_fun U T t2)
+      | _ =>
+        if has_call t then
+          (case try HOLogic.dest_prodT U of
+            SOME (U1, U2) => if U1 = T then raw_massage_fun T U2 t else invalid_map ctxt t
+          | NONE => invalid_map ctxt t)
+        else
+          mk_comp bound_Ts (t, build_map_fst (U, T)));
+
+    fun massage_map (Type (_, Us)) (Type (s, Ts)) t =
+        (case try (dest_map ctxt s) t of
+          SOME (map0, fs) =>
+          let
+            val Type (_, ran_Ts) = range_type (typof t);
+            val map' = mk_map (length fs) Us ran_Ts map0;
+            val fs' = map_flattened_map_args ctxt s (map3 massage_map_or_map_arg Us Ts) fs;
+          in
+            Term.list_comb (map', fs')
+          end
+        | NONE => raise AINT_NO_MAP t)
+      | massage_map _ _ t = raise AINT_NO_MAP t
+    and massage_map_or_map_arg U T t =
+      if T = U then
+        tap check_no_call t
+      else
+        massage_map U T t
+        handle AINT_NO_MAP _ => massage_mutual_fun U T t;
+
+    fun massage_call (t as t1 $ t2) =
+        if has_call t then
+          if t2 = y then
+            massage_map yU yT (elim_y t1) $ y'
+            handle AINT_NO_MAP t' => invalid_map ctxt t'
+          else
+            let val (g, xs) = Term.strip_comb t2 in
+              if g = y then
+                if exists has_call xs then unexpected_rec_call ctxt t2
+                else Term.list_comb (massage_call (mk_compN (length xs) bound_Ts (t1, y)), xs)
+              else
+                ill_formed_rec_call ctxt t
+            end
+        else
+          elim_y t
+      | massage_call t = if t = y then y_of_y' () else ill_formed_rec_call ctxt t;
+  in
+    massage_call
+  end;
+
+fun rec_specs_of bs arg_Ts res_Ts get_indices callssss0 lthy =
+  let
+    val thy = Proof_Context.theory_of lthy;
+
+    val ((missing_arg_Ts, perm0_kks,
+          fp_sugars as {nested_bnfs, fp_res = {xtor_co_iterss = ctor_iters1 :: _, ...},
+            co_inducts = [induct_thm], ...} :: _, (lfp_sugar_thms, _)), lthy') =
+      nested_to_mutual_fps Least_FP bs arg_Ts get_indices callssss0 lthy;
+
+    val perm_fp_sugars = sort (int_ord o pairself #index) fp_sugars;
+
+    val indices = map #index fp_sugars;
+    val perm_indices = map #index perm_fp_sugars;
+
+    val perm_ctrss = map (#ctrs o of_fp_sugar #ctr_sugars) perm_fp_sugars;
+    val perm_ctr_Tsss = map (map (binder_types o fastype_of)) perm_ctrss;
+    val perm_lfpTs = map (body_type o fastype_of o hd) perm_ctrss;
+
+    val nn0 = length arg_Ts;
+    val nn = length perm_lfpTs;
+    val kks = 0 upto nn - 1;
+    val perm_ns = map length perm_ctr_Tsss;
+    val perm_mss = map (map length) perm_ctr_Tsss;
+
+    val perm_Cs = map (body_type o fastype_of o co_rec_of o of_fp_sugar (#xtor_co_iterss o #fp_res))
+      perm_fp_sugars;
+    val perm_fun_arg_Tssss =
+      mk_iter_fun_arg_types perm_ctr_Tsss perm_ns perm_mss (co_rec_of ctor_iters1);
+
+    fun unpermute0 perm0_xs = permute_like (op =) perm0_kks kks perm0_xs;
+    fun unpermute perm_xs = permute_like (op =) perm_indices indices perm_xs;
+
+    val induct_thms = unpermute0 (conj_dests nn induct_thm);
+
+    val lfpTs = unpermute perm_lfpTs;
+    val Cs = unpermute perm_Cs;
+
+    val As_rho = tvar_subst thy (take nn0 lfpTs) arg_Ts;
+    val Cs_rho = map (fst o dest_TVar) Cs ~~ pad_list HOLogic.unitT nn res_Ts;
+
+    val substA = Term.subst_TVars As_rho;
+    val substAT = Term.typ_subst_TVars As_rho;
+    val substCT = Term.typ_subst_TVars Cs_rho;
+    val substACT = substAT o substCT;
+
+    val perm_Cs' = map substCT perm_Cs;
+
+    fun offset_of_ctr 0 _ = 0
+      | offset_of_ctr n (({ctrs, ...} : ctr_sugar) :: ctr_sugars) =
+        length ctrs + offset_of_ctr (n - 1) ctr_sugars;
+
+    fun call_of [i] [T] = (if exists_subtype_in Cs T then Nested_Rec else No_Rec) (i, substACT T)
+      | call_of [i, i'] [T, T'] = Mutual_Rec ((i, substACT T), (i', substACT T'));
+
+    fun mk_ctr_spec ctr offset fun_arg_Tss rec_thm =
+      let
+        val (fun_arg_hss, _) = indexedd fun_arg_Tss 0;
+        val fun_arg_hs = flat_rec_arg_args fun_arg_hss;
+        val fun_arg_iss = map (map (find_index_eq fun_arg_hs)) fun_arg_hss;
+      in
+        {ctr = substA ctr, offset = offset, calls = map2 call_of fun_arg_iss fun_arg_Tss,
+         rec_thm = rec_thm}
+      end;
+
+    fun mk_ctr_specs index (ctr_sugars : ctr_sugar list) iter_thmsss =
+      let
+        val ctrs = #ctrs (nth ctr_sugars index);
+        val rec_thms = co_rec_of (nth iter_thmsss index);
+        val k = offset_of_ctr index ctr_sugars;
+        val n = length ctrs;
+      in
+        map4 mk_ctr_spec ctrs (k upto k + n - 1) (nth perm_fun_arg_Tssss index) rec_thms
+      end;
+
+    fun mk_spec ({T, index, ctr_sugars, co_iterss = iterss, co_iter_thmsss = iter_thmsss, ...}
+      : fp_sugar) =
+      {recx = mk_co_iter thy Least_FP (substAT T) perm_Cs' (co_rec_of (nth iterss index)),
+       nested_map_idents = map (unfold_thms lthy @{thms id_def} o map_id0_of_bnf) nested_bnfs,
+       nested_map_comps = map map_comp_of_bnf nested_bnfs,
+       ctr_specs = mk_ctr_specs index ctr_sugars iter_thmsss};
+  in
+    ((is_some lfp_sugar_thms, map mk_spec fp_sugars, missing_arg_Ts, induct_thm, induct_thms),
+     lthy')
+  end;
+
+val undef_const = Const (@{const_name undefined}, dummyT);
+
+fun permute_args n t =
+  list_comb (t, map Bound (0 :: (n downto 1))) |> fold (K (Term.abs (Name.uu, dummyT))) (0 upto n);
+
+type eqn_data = {
+  fun_name: string,
+  rec_type: typ,
+  ctr: term,
+  ctr_args: term list,
+  left_args: term list,
+  right_args: term list,
+  res_type: typ,
+  rhs_term: term,
+  user_eqn: term
+};
+
+fun dissect_eqn lthy fun_names eqn' =
+  let
+    val eqn = drop_All eqn' |> HOLogic.dest_Trueprop
+      handle TERM _ =>
+        primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
+    val (lhs, rhs) = HOLogic.dest_eq eqn
+        handle TERM _ =>
+          primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
+    val (fun_name, args) = strip_comb lhs
+      |>> (fn x => if is_Free x then fst (dest_Free x)
+          else primrec_error_eqn "malformed function equation (does not start with free)" eqn);
+    val (left_args, rest) = take_prefix is_Free args;
+    val (nonfrees, right_args) = take_suffix is_Free rest;
+    val num_nonfrees = length nonfrees;
+    val _ = num_nonfrees = 1 orelse if num_nonfrees = 0 then
+      primrec_error_eqn "constructor pattern missing in left-hand side" eqn else
+      primrec_error_eqn "more than one non-variable argument in left-hand side" eqn;
+    val _ = member (op =) fun_names fun_name orelse
+      primrec_error_eqn "malformed function equation (does not start with function name)" eqn
+
+    val (ctr, ctr_args) = strip_comb (the_single nonfrees);
+    val _ = try (num_binder_types o fastype_of) ctr = SOME (length ctr_args) orelse
+      primrec_error_eqn "partially applied constructor in pattern" eqn;
+    val _ = let val d = duplicates (op =) (left_args @ ctr_args @ right_args) in null d orelse
+      primrec_error_eqn ("duplicate variable \"" ^ Syntax.string_of_term lthy (hd d) ^
+        "\" in left-hand side") eqn end;
+    val _ = forall is_Free ctr_args orelse
+      primrec_error_eqn "non-primitive pattern in left-hand side" eqn;
+    val _ =
+      let val b = fold_aterms (fn x as Free (v, _) =>
+        if (not (member (op =) (left_args @ ctr_args @ right_args) x) andalso
+        not (member (op =) fun_names v) andalso
+        not (Variable.is_fixed lthy v)) then cons x else I | _ => I) rhs []
+      in
+        null b orelse
+        primrec_error_eqn ("extra variable(s) in right-hand side: " ^
+          commas (map (Syntax.string_of_term lthy) b)) eqn
+      end;
+  in
+    {fun_name = fun_name,
+     rec_type = body_type (type_of ctr),
+     ctr = ctr,
+     ctr_args = ctr_args,
+     left_args = left_args,
+     right_args = right_args,
+     res_type = map fastype_of (left_args @ right_args) ---> fastype_of rhs,
+     rhs_term = rhs,
+     user_eqn = eqn'}
+  end;
+
+fun rewrite_map_arg get_ctr_pos rec_type res_type =
+  let
+    val pT = HOLogic.mk_prodT (rec_type, res_type);
+
+    val maybe_suc = Option.map (fn x => x + 1);
+    fun subst d (t as Bound d') = t |> d = SOME d' ? curry (op $) (fst_const pT)
+      | subst d (Abs (v, T, b)) = Abs (v, if d = SOME ~1 then pT else T, subst (maybe_suc d) b)
+      | subst d t =
+        let
+          val (u, vs) = strip_comb t;
+          val ctr_pos = try (get_ctr_pos o fst o dest_Free) u |> the_default ~1;
+        in
+          if ctr_pos >= 0 then
+            if d = SOME ~1 andalso length vs = ctr_pos then
+              list_comb (permute_args ctr_pos (snd_const pT), vs)
+            else if length vs > ctr_pos andalso is_some d
+                andalso d = try (fn Bound n => n) (nth vs ctr_pos) then
+              list_comb (snd_const pT $ nth vs ctr_pos, map (subst d) (nth_drop ctr_pos vs))
+            else
+              primrec_error_eqn ("recursive call not directly applied to constructor argument") t
+          else
+            list_comb (u, map (subst (d |> d = SOME ~1 ? K NONE)) vs)
+        end
+  in
+    subst (SOME ~1)
+  end;
+
+fun subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls =
+  let
+    fun try_nested_rec bound_Ts y t =
+      AList.lookup (op =) nested_calls y
+      |> Option.map (fn y' =>
+        massage_nested_rec_call lthy has_call (rewrite_map_arg get_ctr_pos) bound_Ts y y' t);
+
+    fun subst bound_Ts (t as g' $ y) =
+        let
+          fun subst_rec () = subst bound_Ts g' $ subst bound_Ts y;
+          val y_head = head_of y;
+        in
+          if not (member (op =) ctr_args y_head) then
+            subst_rec ()
+          else
+            (case try_nested_rec bound_Ts y_head t of
+              SOME t' => t'
+            | NONE =>
+              let val (g, g_args) = strip_comb g' in
+                (case try (get_ctr_pos o fst o dest_Free) g of
+                  SOME ctr_pos =>
+                  (length g_args >= ctr_pos orelse
+                   primrec_error_eqn "too few arguments in recursive call" t;
+                   (case AList.lookup (op =) mutual_calls y of
+                     SOME y' => list_comb (y', g_args)
+                   | NONE => subst_rec ()))
+                | NONE => subst_rec ())
+              end)
+        end
+      | subst bound_Ts (Abs (v, T, b)) = Abs (v, T, subst (T :: bound_Ts) b)
+      | subst _ t = t
+
+    fun subst' t =
+      if has_call t then
+        (* FIXME detect this case earlier? *)
+        primrec_error_eqn "recursive call not directly applied to constructor argument" t
+      else
+        try_nested_rec [] (head_of t) t |> the_default t
+  in
+    subst' o subst []
+  end;
+
+fun build_rec_arg lthy (funs_data : eqn_data list list) has_call (ctr_spec : rec_ctr_spec)
+    (maybe_eqn_data : eqn_data option) =
+  (case maybe_eqn_data of
+    NONE => undef_const
+  | SOME {ctr_args, left_args, right_args, rhs_term = t, ...} =>
+    let
+      val calls = #calls ctr_spec;
+      val n_args = fold (Integer.add o (fn Mutual_Rec _ => 2 | _ => 1)) calls 0;
+
+      val no_calls' = tag_list 0 calls
+        |> map_filter (try (apsnd (fn No_Rec p => p | Mutual_Rec (p, _) => p)));
+      val mutual_calls' = tag_list 0 calls
+        |> map_filter (try (apsnd (fn Mutual_Rec (_, p) => p)));
+      val nested_calls' = tag_list 0 calls
+        |> map_filter (try (apsnd (fn Nested_Rec p => p)));
+
+      val args = replicate n_args ("", dummyT)
+        |> Term.rename_wrt_term t
+        |> map Free
+        |> fold (fn (ctr_arg_idx, (arg_idx, _)) =>
+            nth_map arg_idx (K (nth ctr_args ctr_arg_idx)))
+          no_calls'
+        |> fold (fn (ctr_arg_idx, (arg_idx, T)) =>
+            nth_map arg_idx (K (retype_free T (nth ctr_args ctr_arg_idx))))
+          mutual_calls'
+        |> fold (fn (ctr_arg_idx, (arg_idx, T)) =>
+            nth_map arg_idx (K (retype_free T (nth ctr_args ctr_arg_idx))))
+          nested_calls';
+
+      val fun_name_ctr_pos_list =
+        map (fn (x :: _) => (#fun_name x, length (#left_args x))) funs_data;
+      val get_ctr_pos = try (the o AList.lookup (op =) fun_name_ctr_pos_list) #> the_default ~1;
+      val mutual_calls = map (apfst (nth ctr_args) o apsnd (nth args o fst)) mutual_calls';
+      val nested_calls = map (apfst (nth ctr_args) o apsnd (nth args o fst)) nested_calls';
+    in
+      t
+      |> subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls
+      |> fold_rev lambda (args @ left_args @ right_args)
+    end);
+
+fun build_defs lthy bs mxs (funs_data : eqn_data list list) (rec_specs : rec_spec list) has_call =
+  let
+    val n_funs = length funs_data;
+
+    val ctr_spec_eqn_data_list' =
+      (take n_funs rec_specs |> map #ctr_specs) ~~ funs_data
+      |> maps (uncurry (finds (fn (x, y) => #ctr x = #ctr y))
+          ##> (fn x => null x orelse
+            primrec_error_eqns "excess equations in definition" (map #rhs_term x)) #> fst);
+    val _ = ctr_spec_eqn_data_list' |> map (fn (_, x) => length x <= 1 orelse
+      primrec_error_eqns ("multiple equations for constructor") (map #user_eqn x));
+
+    val ctr_spec_eqn_data_list =
+      ctr_spec_eqn_data_list' @ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
+
+    val recs = take n_funs rec_specs |> map #recx;
+    val rec_args = ctr_spec_eqn_data_list
+      |> sort ((op <) o pairself (#offset o fst) |> make_ord)
+      |> map (uncurry (build_rec_arg lthy funs_data has_call) o apsnd (try the_single));
+    val ctr_poss = map (fn x =>
+      if length (distinct ((op =) o pairself (length o #left_args)) x) <> 1 then
+        primrec_error ("inconstant constructor pattern position for function " ^
+          quote (#fun_name (hd x)))
+      else
+        hd x |> #left_args |> length) funs_data;
+  in
+    (recs, ctr_poss)
+    |-> map2 (fn recx => fn ctr_pos => list_comb (recx, rec_args) |> permute_args ctr_pos)
+    |> Syntax.check_terms lthy
+    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
+      bs mxs
+  end;
+
+fun find_rec_calls has_call ({ctr, ctr_args, rhs_term, ...} : eqn_data) =
+  let
+    fun find bound_Ts (Abs (_, T, b)) ctr_arg = find (T :: bound_Ts) b ctr_arg
+      | find bound_Ts (t as _ $ _) ctr_arg =
+        let
+          val typof = curry fastype_of1 bound_Ts;
+          val (f', args') = strip_comb t;
+          val n = find_index (equal ctr_arg o head_of) args';
+        in
+          if n < 0 then
+            find bound_Ts f' ctr_arg @ maps (fn x => find bound_Ts x ctr_arg) args'
+          else
+            let
+              val (f, args as arg :: _) = chop n args' |>> curry list_comb f'
+              val (arg_head, arg_args) = Term.strip_comb arg;
+            in
+              if has_call f then
+                mk_partial_compN (length arg_args) (typof arg_head) f ::
+                maps (fn x => find bound_Ts x ctr_arg) args
+              else
+                find bound_Ts f ctr_arg @ maps (fn x => find bound_Ts x ctr_arg) args
+            end
+        end
+      | find _ _ _ = [];
+  in
+    map (find [] rhs_term) ctr_args
+    |> (fn [] => NONE | callss => SOME (ctr, callss))
+  end;
+
+fun mk_primrec_tac ctxt num_extra_args map_idents map_comps fun_defs recx =
+  unfold_thms_tac ctxt fun_defs THEN
+  HEADGOAL (rtac (funpow num_extra_args (fn thm => thm RS fun_cong) recx RS trans)) THEN
+  unfold_thms_tac ctxt (@{thms id_def split o_def fst_conv snd_conv} @ map_comps @ map_idents) THEN
+  HEADGOAL (rtac refl);
+
+fun prepare_primrec fixes specs lthy =
+  let
+    val thy = Proof_Context.theory_of lthy;
+
+    val (bs, mxs) = map_split (apfst fst) fixes;
+    val fun_names = map Binding.name_of bs;
+    val eqns_data = map (dissect_eqn lthy fun_names) specs;
+    val funs_data = eqns_data
+      |> partition_eq ((op =) o pairself #fun_name)
+      |> finds (fn (x, y) => x = #fun_name (hd y)) fun_names |> fst
+      |> map (fn (x, y) => the_single y handle List.Empty =>
+          primrec_error ("missing equations for function " ^ quote x));
+
+    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
+    val arg_Ts = map (#rec_type o hd) funs_data;
+    val res_Ts = map (#res_type o hd) funs_data;
+    val callssss = funs_data
+      |> map (partition_eq ((op =) o pairself #ctr))
+      |> map (maps (map_filter (find_rec_calls has_call)));
+
+    val _ = (case filter_out (fn (_, T) => Sign.of_sort thy (T, HOLogic.typeS)) (bs ~~ res_Ts) of
+        [] => ()
+      | (b, _) :: _ => primrec_error ("type of " ^ Binding.print b ^ " contains top sort"));
+
+    val ((n2m, rec_specs, _, induct_thm, induct_thms), lthy') =
+      rec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
+
+    val actual_nn = length funs_data;
+
+    val _ = let val ctrs = (maps (map #ctr o #ctr_specs) rec_specs) in
+      map (fn {ctr, user_eqn, ...} => member (op =) ctrs ctr orelse
+        primrec_error_eqn ("argument " ^ quote (Syntax.string_of_term lthy' ctr) ^
+          " is not a constructor in left-hand side") user_eqn) eqns_data end;
+
+    val defs = build_defs lthy' bs mxs funs_data rec_specs has_call;
+
+    fun prove lthy def_thms' ({ctr_specs, nested_map_idents, nested_map_comps, ...} : rec_spec)
+        (fun_data : eqn_data list) =
+      let
+        val def_thms = map (snd o snd) def_thms';
+        val simp_thmss = finds (fn (x, y) => #ctr x = #ctr y) fun_data ctr_specs
+          |> fst
+          |> map_filter (try (fn (x, [y]) =>
+            (#user_eqn x, length (#left_args x) + length (#right_args x), #rec_thm y)))
+          |> map (fn (user_eqn, num_extra_args, rec_thm) =>
+            mk_primrec_tac lthy num_extra_args nested_map_idents nested_map_comps def_thms rec_thm
+            |> K |> Goal.prove lthy [] [] user_eqn
+            |> Thm.close_derivation);
+        val poss = find_indices (fn (x, y) => #ctr x = #ctr y) fun_data eqns_data;
+      in
+        (poss, simp_thmss)
+      end;
+
+    val notes =
+      (if n2m then map2 (fn name => fn thm =>
+        (name, inductN, [thm], [])) fun_names (take actual_nn induct_thms) else [])
+      |> map (fn (prefix, thmN, thms, attrs) =>
+        ((Binding.qualify true prefix (Binding.name thmN), attrs), [(thms, [])]));
+
+    val common_name = mk_common_name fun_names;
+
+    val common_notes =
+      (if n2m then [(inductN, [induct_thm], [])] else [])
+      |> map (fn (thmN, thms, attrs) =>
+        ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
+  in
+    (((fun_names, defs),
+      fn lthy => fn defs =>
+        split_list (map2 (prove lthy defs) (take actual_nn rec_specs) funs_data)),
+      lthy' |> Local_Theory.notes (notes @ common_notes) |> snd)
+  end;
+
+(* primrec definition *)
+
+fun add_primrec_simple fixes ts lthy =
+  let
+    val (((names, defs), prove), lthy) = prepare_primrec fixes ts lthy
+      handle ERROR str => primrec_error str;
+  in
+    lthy
+    |> fold_map Local_Theory.define defs
+    |-> (fn defs => `(fn lthy => (names, (map fst defs, prove lthy defs))))
+  end
+  handle Primrec_Error (str, eqns) =>
+    if null eqns
+    then error ("primrec_new error:\n  " ^ str)
+    else error ("primrec_new error:\n  " ^ str ^ "\nin\n  " ^
+      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
+
+local
+
+fun gen_primrec prep_spec (raw_fixes : (binding * 'a option * mixfix) list) raw_spec lthy =
+  let
+    val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes)
+    val _ = null d orelse primrec_error ("duplicate function name(s): " ^ commas d);
+
+    val (fixes, specs) = fst (prep_spec raw_fixes raw_spec lthy);
+
+    val mk_notes =
+      flat ooo map3 (fn poss => fn prefix => fn thms =>
+        let
+          val (bs, attrss) = map_split (fst o nth specs) poss;
+          val notes =
+            map3 (fn b => fn attrs => fn thm =>
+              ((Binding.qualify false prefix b, code_nitpicksimp_simp_attrs @ attrs), [([thm], [])]))
+            bs attrss thms;
+        in
+          ((Binding.qualify true prefix (Binding.name simpsN), []), [(thms, [])]) :: notes
+        end);
+  in
+    lthy
+    |> add_primrec_simple fixes (map snd specs)
+    |-> (fn (names, (ts, (posss, simpss))) =>
+      Spec_Rules.add Spec_Rules.Equational (ts, flat simpss)
+      #> Local_Theory.notes (mk_notes posss names simpss)
+      #>> pair ts o map snd)
+  end;
+
+in
+
+val add_primrec = gen_primrec Specification.check_spec;
+val add_primrec_cmd = gen_primrec Specification.read_spec;
+
+end;
+
+fun add_primrec_global fixes specs thy =
+  let
+    val lthy = Named_Target.theory_init thy;
+    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
+    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
+  in ((ts, simps'), Local_Theory.exit_global lthy') end;
+
+fun add_primrec_overloaded ops fixes specs thy =
+  let
+    val lthy = Overloading.overloading ops thy;
+    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
+    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
+  in ((ts, simps'), Local_Theory.exit_global lthy') end;
+
+end;
--- a/src/HOL/BNF/Tools/ctr_sugar.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/BNF/Tools/ctr_sugar.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -33,6 +33,7 @@
      case_conv_ifs: thm list};
 
   val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar
+  val transfer_ctr_sugar: Proof.context -> ctr_sugar -> ctr_sugar
   val ctr_sugar_of: Proof.context -> string -> ctr_sugar option
   val ctr_sugars_of: Proof.context -> ctr_sugar list
 
@@ -174,10 +175,11 @@
 val dest_attrs = @{attributes [dest]};
 val safe_elim_attrs = @{attributes [elim!]};
 val iff_attrs = @{attributes [iff]};
-val induct_simp_attrs = @{attributes [induct_simp]};
-val nitpick_attrs = @{attributes [nitpick_simp]};
+val inductsimp_attrs = @{attributes [induct_simp]};
+val nitpicksimp_attrs = @{attributes [nitpick_simp]};
 val simp_attrs = @{attributes [simp]};
-val code_nitpick_simp_simp_attrs = Code.add_default_eqn_attrib :: nitpick_attrs @ simp_attrs;
+val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
+val code_nitpicksimp_simp_attrs = code_nitpicksimp_attrs @ simp_attrs;
 
 fun unflat_lookup eq xs ys = map (fn xs' => permute_like eq xs xs' ys);
 
@@ -391,7 +393,8 @@
          Term.lambda w (Library.foldr1 HOLogic.mk_disj (map3 mk_case_disj xctrs xfs xss)));
 
     val ((raw_case, (_, raw_case_def)), (lthy', lthy)) = no_defs_lthy
-      |> Local_Theory.define ((case_binding, NoSyn), ((Thm.def_binding case_binding, []), case_rhs))
+      |> Local_Theory.define ((case_binding, NoSyn),
+        ((Binding.conceal (Thm.def_binding case_binding), []), case_rhs))
       ||> `Local_Theory.restore;
 
     val phi = Proof_Context.export_morphism lthy lthy';
@@ -869,8 +872,15 @@
         val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases));
         val cases_type_attr = Attrib.internal (K (Induct.cases_type fcT_name));
 
+        val anonymous_notes =
+          [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs),
+           (map (fn th => th RS @{thm eq_False[THEN iffD2]}
+              handle THM _ => th RS @{thm eq_True[THEN iffD2]}) nontriv_disc_thms,
+            code_nitpicksimp_attrs)]
+          |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
+
         val notes =
-          [(caseN, case_thms, code_nitpick_simp_simp_attrs),
+          [(caseN, case_thms, code_nitpicksimp_simp_attrs),
            (case_congN, [case_cong_thm], []),
            (case_conv_ifN, case_conv_if_thms, []),
            (collapseN, safe_collapse_thms, simp_attrs),
@@ -878,12 +888,12 @@
            (discIN, nontriv_discI_thms, []),
            (disc_excludeN, disc_exclude_thms, dest_attrs),
            (disc_exhaustN, disc_exhaust_thms, [exhaust_case_names_attr]),
-           (distinctN, distinct_thms, simp_attrs @ induct_simp_attrs),
+           (distinctN, distinct_thms, simp_attrs @ inductsimp_attrs),
            (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]),
            (expandN, expand_thms, []),
-           (injectN, inject_thms, iff_attrs @ induct_simp_attrs),
+           (injectN, inject_thms, iff_attrs @ inductsimp_attrs),
            (nchotomyN, [nchotomy_thm], []),
-           (selN, all_sel_thms, code_nitpick_simp_simp_attrs),
+           (selN, all_sel_thms, code_nitpicksimp_simp_attrs),
            (sel_exhaustN, sel_exhaust_thms, [exhaust_case_names_attr]),
            (sel_splitN, sel_split_thms, []),
            (sel_split_asmN, sel_split_asm_thms, []),
@@ -895,10 +905,6 @@
           |> map (fn (thmN, thms, attrs) =>
             ((qualify true (Binding.name thmN), attrs), [(thms, [])]));
 
-        val notes' =
-          [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs)]
-          |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
-
         val ctr_sugar =
           {ctrs = ctrs, casex = casex, discs = discs, selss = selss, exhaust = exhaust_thm,
            nchotomy = nchotomy_thm, injects = inject_thms, distincts = distinct_thms,
@@ -915,7 +921,7 @@
             (Local_Theory.declaration {syntax = false, pervasive = true}
                (fn phi => Case_Translation.register
                   (Morphism.term phi casex) (map (Morphism.term phi) ctrs)))
-         |> Local_Theory.notes (notes' @ notes) |> snd
+         |> Local_Theory.notes (anonymous_notes @ notes) |> snd
          |> register_ctr_sugar fcT_name ctr_sugar)
       end;
   in
--- a/src/HOL/Big_Operators.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Big_Operators.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -696,11 +696,7 @@
 lemma setsum_subtractf:
   "setsum (%x. ((f x)::'a::ab_group_add) - g x) A =
     setsum f A - setsum g A"
-proof (cases "finite A")
-  case True thus ?thesis by (simp add: diff_minus setsum_addf setsum_negf)
-next
-  case False thus ?thesis by simp
-qed
+  using setsum_addf [of f "- g" A] by (simp add: setsum_negf)
 
 lemma setsum_nonneg:
   assumes nn: "\<forall>x\<in>A. (0::'a::{ordered_ab_semigroup_add,comm_monoid_add}) \<le> f x"
@@ -1999,35 +1995,35 @@
   assumes fin_nonempty: "finite A" "A \<noteq> {}"
 begin
 
-lemma Min_ge_iff [simp, no_atp]:
+lemma Min_ge_iff [simp]:
   "x \<le> Min A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
   using fin_nonempty by (fact Min.bounded_iff)
 
-lemma Max_le_iff [simp, no_atp]:
+lemma Max_le_iff [simp]:
   "Max A \<le> x \<longleftrightarrow> (\<forall>a\<in>A. a \<le> x)"
   using fin_nonempty by (fact Max.bounded_iff)
 
-lemma Min_gr_iff [simp, no_atp]:
+lemma Min_gr_iff [simp]:
   "x < Min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
   using fin_nonempty  by (induct rule: finite_ne_induct) simp_all
 
-lemma Max_less_iff [simp, no_atp]:
+lemma Max_less_iff [simp]:
   "Max A < x \<longleftrightarrow> (\<forall>a\<in>A. a < x)"
   using fin_nonempty by (induct rule: finite_ne_induct) simp_all
 
-lemma Min_le_iff [no_atp]:
+lemma Min_le_iff:
   "Min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
   using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: min_le_iff_disj)
 
-lemma Max_ge_iff [no_atp]:
+lemma Max_ge_iff:
   "x \<le> Max A \<longleftrightarrow> (\<exists>a\<in>A. x \<le> a)"
   using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: le_max_iff_disj)
 
-lemma Min_less_iff [no_atp]:
+lemma Min_less_iff:
   "Min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
   using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: min_less_iff_disj)
 
-lemma Max_gr_iff [no_atp]:
+lemma Max_gr_iff:
   "x < Max A \<longleftrightarrow> (\<exists>a\<in>A. x < a)"
   using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: less_max_iff_disj)
 
--- a/src/HOL/Codegenerator_Test/Generate_Binary_Nat.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Codegenerator_Test/Generate_Binary_Nat.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -16,6 +16,20 @@
   by a corresponding @{text export_code} command.
 *}
 
-export_code _ checking SML OCaml? Haskell? Scala
+text {* Formal joining of hierarchy of implicit definitions in Scala *}
+
+class semiring_numeral_even_odd = semiring_numeral_div + even_odd
+
+instance nat :: semiring_numeral_even_odd ..
+
+definition semiring_numeral_even_odd :: "'a itself \<Rightarrow> 'a::semiring_numeral_even_odd"
+where
+  "semiring_numeral_even_odd TYPE('a) = undefined"
+
+definition semiring_numeral_even_odd_nat :: "nat itself \<Rightarrow> nat"
+where
+  "semiring_numeral_even_odd_nat = semiring_numeral_even_odd"
+
+export_code _ checking  SML OCaml? Haskell? Scala
 
 end
--- a/src/HOL/Codegenerator_Test/Generate_Efficient_Datastructures.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Codegenerator_Test/Generate_Efficient_Datastructures.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -26,7 +26,7 @@
   "pred_of_set = pred_of_set" ..
 
 lemma [code, code del]:
-  "acc = acc" ..
+  "Wellfounded.acc = Wellfounded.acc" ..
 
 lemma [code, code del]:
   "Cardinality.card' = Cardinality.card'" ..
--- a/src/HOL/Complete_Lattices.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Complete_Lattices.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -15,10 +15,66 @@
 
 class Inf =
   fixes Inf :: "'a set \<Rightarrow> 'a" ("\<Sqinter>_" [900] 900)
+begin
+
+definition INFI :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
+  INF_def: "INFI A f = \<Sqinter>(f ` A)"
+
+lemma INF_image [simp]: "INFI (f`A) g = INFI A (\<lambda>x. g (f x))"
+  by (simp add: INF_def image_image)
+
+lemma INF_cong: "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> INFI A C = INFI B D"
+  by (simp add: INF_def image_def)
+
+end
 
 class Sup =
   fixes Sup :: "'a set \<Rightarrow> 'a" ("\<Squnion>_" [900] 900)
+begin
 
+definition SUPR :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
+  SUP_def: "SUPR A f = \<Squnion>(f ` A)"
+
+lemma SUP_image [simp]: "SUPR (f`A) g = SUPR A (%x. g (f x))"
+  by (simp add: SUP_def image_image)
+
+lemma SUP_cong: "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> SUPR A C = SUPR B D"
+  by (simp add: SUP_def image_def)
+
+end
+
+text {*
+  Note: must use names @{const INFI} and @{const SUPR} here instead of
+  @{text INF} and @{text SUP} to allow the following syntax coexist
+  with the plain constant names.
+*}
+
+syntax
+  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3INF _./ _)" [0, 10] 10)
+  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3INF _:_./ _)" [0, 0, 10] 10)
+  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3SUP _./ _)" [0, 10] 10)
+  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3SUP _:_./ _)" [0, 0, 10] 10)
+
+syntax (xsymbols)
+  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
+  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
+  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
+  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
+
+translations
+  "INF x y. B"   == "INF x. INF y. B"
+  "INF x. B"     == "CONST INFI CONST UNIV (%x. B)"
+  "INF x. B"     == "INF x:CONST UNIV. B"
+  "INF x:A. B"   == "CONST INFI A (%x. B)"
+  "SUP x y. B"   == "SUP x. SUP y. B"
+  "SUP x. B"     == "CONST SUPR CONST UNIV (%x. B)"
+  "SUP x. B"     == "SUP x:CONST UNIV. B"
+  "SUP x:A. B"   == "CONST SUPR A (%x. B)"
+
+print_translation {*
+  [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INFI} @{syntax_const "_INF"},
+    Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax SUPR} @{syntax_const "_SUP"}]
+*} -- {* to avoid eta-contraction of body *}
 
 subsection {* Abstract complete lattices *}
 
@@ -49,59 +105,17 @@
     (unfold_locales, (fact Inf_empty Sup_empty
         Sup_upper Sup_least Inf_lower Inf_greatest)+)
 
-definition INFI :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
-  INF_def: "INFI A f = \<Sqinter>(f ` A)"
-
-definition SUPR :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
-  SUP_def: "SUPR A f = \<Squnion>(f ` A)"
-
-text {*
-  Note: must use names @{const INFI} and @{const SUPR} here instead of
-  @{text INF} and @{text SUP} to allow the following syntax coexist
-  with the plain constant names.
-*}
-
 end
 
-syntax
-  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3INF _./ _)" [0, 10] 10)
-  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3INF _:_./ _)" [0, 0, 10] 10)
-  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3SUP _./ _)" [0, 10] 10)
-  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3SUP _:_./ _)" [0, 0, 10] 10)
-
-syntax (xsymbols)
-  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
-  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
-  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
-  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
-
-translations
-  "INF x y. B"   == "INF x. INF y. B"
-  "INF x. B"     == "CONST INFI CONST UNIV (%x. B)"
-  "INF x. B"     == "INF x:CONST UNIV. B"
-  "INF x:A. B"   == "CONST INFI A (%x. B)"
-  "SUP x y. B"   == "SUP x. SUP y. B"
-  "SUP x. B"     == "CONST SUPR CONST UNIV (%x. B)"
-  "SUP x. B"     == "SUP x:CONST UNIV. B"
-  "SUP x:A. B"   == "CONST SUPR A (%x. B)"
-
-print_translation {*
-  [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INFI} @{syntax_const "_INF"},
-    Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax SUPR} @{syntax_const "_SUP"}]
-*} -- {* to avoid eta-contraction of body *}
-
 context complete_lattice
 begin
 
-lemma INF_foundation_dual [no_atp]:
-  "complete_lattice.SUPR Inf = INFI"
-  by (simp add: fun_eq_iff INF_def
-    complete_lattice.SUP_def [OF dual_complete_lattice])
+lemma INF_foundation_dual:
+  "Sup.SUPR Inf = INFI"
+  by (simp add: fun_eq_iff INF_def Sup.SUP_def)
 
-lemma SUP_foundation_dual [no_atp]:
-  "complete_lattice.INFI Sup = SUPR"
-  by (simp add: fun_eq_iff SUP_def
-    complete_lattice.INF_def [OF dual_complete_lattice])
+lemma SUP_foundation_dual:
+  "Inf.INFI Sup = SUPR" by (simp add: fun_eq_iff SUP_def Inf.INF_def)
 
 lemma Sup_eqI:
   "(\<And>y. y \<in> A \<Longrightarrow> y \<le> x) \<Longrightarrow> (\<And>y. (\<And>z. z \<in> A \<Longrightarrow> z \<le> y) \<Longrightarrow> x \<le> y) \<Longrightarrow> \<Squnion>A = x"
@@ -181,12 +195,6 @@
   "\<Squnion>UNIV = \<top>"
   by (auto intro!: antisym Sup_upper)
 
-lemma INF_image [simp]: "(\<Sqinter>x\<in>f`A. g x) = (\<Sqinter>x\<in>A. g (f x))"
-  by (simp add: INF_def image_image)
-
-lemma SUP_image [simp]: "(\<Squnion>x\<in>f`A. g x) = (\<Squnion>x\<in>A. g (f x))"
-  by (simp add: SUP_def image_image)
-
 lemma Inf_Sup: "\<Sqinter>A = \<Squnion>{b. \<forall>a \<in> A. b \<sqsubseteq> a}"
   by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least)
 
@@ -199,14 +207,6 @@
 lemma Sup_subset_mono: "A \<subseteq> B \<Longrightarrow> \<Squnion>A \<sqsubseteq> \<Squnion>B"
   by (auto intro: Sup_least Sup_upper)
 
-lemma INF_cong:
-  "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> (\<Sqinter>x\<in>A. C x) = (\<Sqinter>x\<in>B. D x)"
-  by (simp add: INF_def image_def)
-
-lemma SUP_cong:
-  "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> (\<Squnion>x\<in>A. C x) = (\<Squnion>x\<in>B. D x)"
-  by (simp add: SUP_def image_def)
-
 lemma Inf_mono:
   assumes "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<sqsubseteq> b"
   shows "\<Sqinter>A \<sqsubseteq> \<Sqinter>B"
@@ -306,7 +306,7 @@
   show "?R \<le> ?L" by (rule SUP_least) (auto intro: le_supI1 le_supI2 SUP_upper)
 qed
 
-lemma Inf_top_conv [simp, no_atp]:
+lemma Inf_top_conv [simp]:
   "\<Sqinter>A = \<top> \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)"
   "\<top> = \<Sqinter>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)"
 proof -
@@ -333,7 +333,7 @@
  "\<top> = (\<Sqinter>x\<in>A. B x) \<longleftrightarrow> (\<forall>x\<in>A. B x = \<top>)"
   by (auto simp add: INF_def)
 
-lemma Sup_bot_conv [simp, no_atp]:
+lemma Sup_bot_conv [simp]:
   "\<Squnion>A = \<bottom> \<longleftrightarrow> (\<forall>x\<in>A. x = \<bottom>)" (is ?P)
   "\<bottom> = \<Squnion>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<bottom>)" (is ?Q)
   using dual_complete_lattice
@@ -769,7 +769,7 @@
     by (simp add: Inf_set_def image_def)
 qed
 
-lemma Inter_iff [simp,no_atp]: "A \<in> \<Inter>C \<longleftrightarrow> (\<forall>X\<in>C. A \<in> X)"
+lemma Inter_iff [simp]: "A \<in> \<Inter>C \<longleftrightarrow> (\<forall>X\<in>C. A \<in> X)"
   by (unfold Inter_eq) blast
 
 lemma InterI [intro!]: "(\<And>X. X \<in> C \<Longrightarrow> A \<in> X) \<Longrightarrow> A \<in> \<Inter>C"
@@ -814,7 +814,7 @@
 lemma Inter_Un_distrib: "\<Inter>(A \<union> B) = \<Inter>A \<inter> \<Inter>B"
   by (fact Inf_union_distrib)
 
-lemma Inter_UNIV_conv [simp, no_atp]:
+lemma Inter_UNIV_conv [simp]:
   "\<Inter>A = UNIV \<longleftrightarrow> (\<forall>x\<in>A. x = UNIV)"
   "UNIV = \<Inter>A \<longleftrightarrow> (\<forall>x\<in>A. x = UNIV)"
   by (fact Inf_top_conv)+
@@ -952,7 +952,7 @@
     by (simp add: Sup_set_def image_def)
 qed
 
-lemma Union_iff [simp, no_atp]:
+lemma Union_iff [simp]:
   "A \<in> \<Union>C \<longleftrightarrow> (\<exists>X\<in>C. A\<in>X)"
   by (unfold Union_eq) blast
 
@@ -987,10 +987,10 @@
 lemma Union_Int_subset: "\<Union>(A \<inter> B) \<subseteq> \<Union>A \<inter> \<Union>B"
   by (fact Sup_inter_less_eq)
 
-lemma Union_empty_conv [no_atp]: "(\<Union>A = {}) \<longleftrightarrow> (\<forall>x\<in>A. x = {})"
+lemma Union_empty_conv: "(\<Union>A = {}) \<longleftrightarrow> (\<forall>x\<in>A. x = {})"
   by (fact Sup_bot_conv) (* already simp *)
 
-lemma empty_Union_conv [no_atp]: "({} = \<Union>A) \<longleftrightarrow> (\<forall>x\<in>A. x = {})"
+lemma empty_Union_conv: "({} = \<Union>A) \<longleftrightarrow> (\<forall>x\<in>A. x = {})"
   by (fact Sup_bot_conv) (* already simp *)
 
 lemma subset_Pow_Union: "A \<subseteq> Pow (\<Union>A)"
@@ -1044,7 +1044,7 @@
   [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax UNION} @{syntax_const "_UNION"}]
 *} -- {* to avoid eta-contraction of body *}
 
-lemma UNION_eq [no_atp]:
+lemma UNION_eq:
   "(\<Union>x\<in>A. B x) = {y. \<exists>x\<in>A. y \<in> B x}"
   by (auto simp add: SUP_def)
 
@@ -1088,13 +1088,13 @@
 lemma UN_least: "(\<And>x. x \<in> A \<Longrightarrow> B x \<subseteq> C) \<Longrightarrow> (\<Union>x\<in>A. B x) \<subseteq> C"
   by (fact SUP_least)
 
-lemma Collect_bex_eq [no_atp]: "{x. \<exists>y\<in>A. P x y} = (\<Union>y\<in>A. {x. P x y})"
+lemma Collect_bex_eq: "{x. \<exists>y\<in>A. P x y} = (\<Union>y\<in>A. {x. P x y})"
   by blast
 
 lemma UN_insert_distrib: "u \<in> A \<Longrightarrow> (\<Union>x\<in>A. insert a (B x)) = insert a (\<Union>x\<in>A. B x)"
   by blast
 
-lemma UN_empty [no_atp]: "(\<Union>x\<in>{}. B x) = {}"
+lemma UN_empty: "(\<Union>x\<in>{}. B x) = {}"
   by (fact SUP_empty)
 
 lemma UN_empty2: "(\<Union>x\<in>A. {}) = {}"
@@ -1126,7 +1126,7 @@
   "(\<Union>x\<in>A. B x) = {} \<longleftrightarrow> (\<forall>x\<in>A. B x = {})"
   by (fact SUP_bot_conv)+ (* already simp *)
 
-lemma Collect_ex_eq [no_atp]: "{x. \<exists>y. P x y} = (\<Union>y. {x. P x y})"
+lemma Collect_ex_eq: "{x. \<exists>y. P x y} = (\<Union>y. {x. P x y})"
   by blast
 
 lemma ball_UN: "(\<forall>z \<in> UNION A B. P z) \<longleftrightarrow> (\<forall>x\<in>A. \<forall>z \<in> B x. P z)"
@@ -1248,7 +1248,7 @@
   "\<And>A B f. (\<Inter>x\<in>f`A. B x) = (\<Inter>a\<in>A. B (f a))"
   by auto
 
-lemma UN_ball_bex_simps [simp, no_atp]:
+lemma UN_ball_bex_simps [simp]:
   "\<And>A P. (\<forall>x\<in>\<Union>A. P x) \<longleftrightarrow> (\<forall>y\<in>A. \<forall>x\<in>y. P x)"
   "\<And>A B P. (\<forall>x\<in>UNION A B. P x) = (\<forall>a\<in>A. \<forall>x\<in> B a. P x)"
   "\<And>A P. (\<exists>x\<in>\<Union>A. P x) \<longleftrightarrow> (\<exists>y\<in>A. \<exists>x\<in>y. P x)"
--- a/src/HOL/Complex.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Complex.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -587,7 +587,7 @@
   by (simp add: cis_def)
 
 lemma cis_divide: "cis a / cis b = cis (a - b)"
-  by (simp add: complex_divide_def cis_mult diff_minus)
+  by (simp add: complex_divide_def cis_mult)
 
 lemma cos_n_Re_cis_pow_n: "cos (real n * a) = Re(cis a ^ n)"
   by (auto simp add: DeMoivre)
--- a/src/HOL/Conditionally_Complete_Lattices.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Conditionally_Complete_Lattices.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1,19 +1,160 @@
 (*  Title:      HOL/Conditionally_Complete_Lattices.thy
     Author:     Amine Chaieb and L C Paulson, University of Cambridge
     Author:     Johannes Hölzl, TU München
+    Author:     Luke S. Serafin, Carnegie Mellon University
 *)
 
 header {* Conditionally-complete Lattices *}
 
 theory Conditionally_Complete_Lattices
-imports Main Lubs
+imports Main
+begin
+
+lemma (in linorder) Sup_fin_eq_Max: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup_fin X = Max X"
+  by (induct X rule: finite_ne_induct) (simp_all add: sup_max)
+
+lemma (in linorder) Inf_fin_eq_Min: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf_fin X = Min X"
+  by (induct X rule: finite_ne_induct) (simp_all add: inf_min)
+
+context preorder
 begin
 
-lemma Sup_fin_eq_Max: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup_fin X = Max X"
-  by (induct X rule: finite_ne_induct) (simp_all add: sup_max)
+definition "bdd_above A \<longleftrightarrow> (\<exists>M. \<forall>x \<in> A. x \<le> M)"
+definition "bdd_below A \<longleftrightarrow> (\<exists>m. \<forall>x \<in> A. m \<le> x)"
+
+lemma bdd_aboveI[intro]: "(\<And>x. x \<in> A \<Longrightarrow> x \<le> M) \<Longrightarrow> bdd_above A"
+  by (auto simp: bdd_above_def)
+
+lemma bdd_belowI[intro]: "(\<And>x. x \<in> A \<Longrightarrow> m \<le> x) \<Longrightarrow> bdd_below A"
+  by (auto simp: bdd_below_def)
+
+lemma bdd_aboveI2: "(\<And>x. x \<in> A \<Longrightarrow> f x \<le> M) \<Longrightarrow> bdd_above (f`A)"
+  by force
+
+lemma bdd_belowI2: "(\<And>x. x \<in> A \<Longrightarrow> m \<le> f x) \<Longrightarrow> bdd_below (f`A)"
+  by force
+
+lemma bdd_above_empty [simp, intro]: "bdd_above {}"
+  unfolding bdd_above_def by auto
+
+lemma bdd_below_empty [simp, intro]: "bdd_below {}"
+  unfolding bdd_below_def by auto
+
+lemma bdd_above_mono: "bdd_above B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> bdd_above A"
+  by (metis (full_types) bdd_above_def order_class.le_neq_trans psubsetD)
+
+lemma bdd_below_mono: "bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> bdd_below A"
+  by (metis bdd_below_def order_class.le_neq_trans psubsetD)
+
+lemma bdd_above_Int1 [simp]: "bdd_above A \<Longrightarrow> bdd_above (A \<inter> B)"
+  using bdd_above_mono by auto
+
+lemma bdd_above_Int2 [simp]: "bdd_above B \<Longrightarrow> bdd_above (A \<inter> B)"
+  using bdd_above_mono by auto
+
+lemma bdd_below_Int1 [simp]: "bdd_below A \<Longrightarrow> bdd_below (A \<inter> B)"
+  using bdd_below_mono by auto
+
+lemma bdd_below_Int2 [simp]: "bdd_below B \<Longrightarrow> bdd_below (A \<inter> B)"
+  using bdd_below_mono by auto
+
+lemma bdd_above_Ioo [simp, intro]: "bdd_above {a <..< b}"
+  by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le)
+
+lemma bdd_above_Ico [simp, intro]: "bdd_above {a ..< b}"
+  by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le)
+
+lemma bdd_above_Iio [simp, intro]: "bdd_above {..< b}"
+  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
+
+lemma bdd_above_Ioc [simp, intro]: "bdd_above {a <.. b}"
+  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
+
+lemma bdd_above_Icc [simp, intro]: "bdd_above {a .. b}"
+  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
+
+lemma bdd_above_Iic [simp, intro]: "bdd_above {.. b}"
+  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
+
+lemma bdd_below_Ioo [simp, intro]: "bdd_below {a <..< b}"
+  by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le)
+
+lemma bdd_below_Ioc [simp, intro]: "bdd_below {a <.. b}"
+  by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le)
+
+lemma bdd_below_Ioi [simp, intro]: "bdd_below {a <..}"
+  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
 
-lemma Inf_fin_eq_Min: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf_fin X = Min X"
-  by (induct X rule: finite_ne_induct) (simp_all add: inf_min)
+lemma bdd_below_Ico [simp, intro]: "bdd_below {a ..< b}"
+  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
+
+lemma bdd_below_Icc [simp, intro]: "bdd_below {a .. b}"
+  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
+
+lemma bdd_below_Ici [simp, intro]: "bdd_below {a ..}"
+  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
+
+end
+
+lemma (in order_top) bdd_above_top[simp, intro!]: "bdd_above A"
+  by (rule bdd_aboveI[of _ top]) simp
+
+lemma (in order_bot) bdd_above_bot[simp, intro!]: "bdd_below A"
+  by (rule bdd_belowI[of _ bot]) simp
+
+lemma bdd_above_uminus[simp]:
+  fixes X :: "'a::ordered_ab_group_add set"
+  shows "bdd_above (uminus ` X) \<longleftrightarrow> bdd_below X"
+  by (auto simp: bdd_above_def bdd_below_def intro: le_imp_neg_le) (metis le_imp_neg_le minus_minus)
+
+lemma bdd_below_uminus[simp]:
+  fixes X :: "'a::ordered_ab_group_add set"
+  shows"bdd_below (uminus ` X) \<longleftrightarrow> bdd_above X"
+  by (auto simp: bdd_above_def bdd_below_def intro: le_imp_neg_le) (metis le_imp_neg_le minus_minus)
+
+context lattice
+begin
+
+lemma bdd_above_insert [simp]: "bdd_above (insert a A) = bdd_above A"
+  by (auto simp: bdd_above_def intro: le_supI2 sup_ge1)
+
+lemma bdd_below_insert [simp]: "bdd_below (insert a A) = bdd_below A"
+  by (auto simp: bdd_below_def intro: le_infI2 inf_le1)
+
+lemma bdd_finite [simp]:
+  assumes "finite A" shows bdd_above_finite: "bdd_above A" and bdd_below_finite: "bdd_below A"
+  using assms by (induct rule: finite_induct, auto)
+
+lemma bdd_above_Un [simp]: "bdd_above (A \<union> B) = (bdd_above A \<and> bdd_above B)"
+proof
+  assume "bdd_above (A \<union> B)"
+  thus "bdd_above A \<and> bdd_above B" unfolding bdd_above_def by auto
+next
+  assume "bdd_above A \<and> bdd_above B"
+  then obtain a b where "\<forall>x\<in>A. x \<le> a" "\<forall>x\<in>B. x \<le> b" unfolding bdd_above_def by auto
+  hence "\<forall>x \<in> A \<union> B. x \<le> sup a b" by (auto intro: Un_iff le_supI1 le_supI2)
+  thus "bdd_above (A \<union> B)" unfolding bdd_above_def ..
+qed
+
+lemma bdd_below_Un [simp]: "bdd_below (A \<union> B) = (bdd_below A \<and> bdd_below B)"
+proof
+  assume "bdd_below (A \<union> B)"
+  thus "bdd_below A \<and> bdd_below B" unfolding bdd_below_def by auto
+next
+  assume "bdd_below A \<and> bdd_below B"
+  then obtain a b where "\<forall>x\<in>A. a \<le> x" "\<forall>x\<in>B. b \<le> x" unfolding bdd_below_def by auto
+  hence "\<forall>x \<in> A \<union> B. inf a b \<le> x" by (auto intro: Un_iff le_infI1 le_infI2)
+  thus "bdd_below (A \<union> B)" unfolding bdd_below_def ..
+qed
+
+lemma bdd_above_sup[simp]: "bdd_above ((\<lambda>x. sup (f x) (g x)) ` A) \<longleftrightarrow> bdd_above (f`A) \<and> bdd_above (g`A)"
+  by (auto simp: bdd_above_def intro: le_supI1 le_supI2)
+
+lemma bdd_below_inf[simp]: "bdd_below ((\<lambda>x. inf (f x) (g x)) ` A) \<longleftrightarrow> bdd_below (f`A) \<and> bdd_below (g`A)"
+  by (auto simp: bdd_below_def intro: le_infI1 le_infI2)
+
+end
+
 
 text {*
 
@@ -23,46 +164,42 @@
 *}
 
 class conditionally_complete_lattice = lattice + Sup + Inf +
-  assumes cInf_lower: "x \<in> X \<Longrightarrow> (\<And>a. a \<in> X \<Longrightarrow> z \<le> a) \<Longrightarrow> Inf X \<le> x"
+  assumes cInf_lower: "x \<in> X \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> x"
     and cInf_greatest: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> Inf X"
-  assumes cSup_upper: "x \<in> X \<Longrightarrow> (\<And>a. a \<in> X \<Longrightarrow> a \<le> z) \<Longrightarrow> x \<le> Sup X"
+  assumes cSup_upper: "x \<in> X \<Longrightarrow> bdd_above X \<Longrightarrow> x \<le> Sup X"
     and cSup_least: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X \<le> z"
 begin
 
-lemma cSup_eq_maximum: (*REAL_SUP_MAX in HOL4*)
-  "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
-  by (blast intro: antisym cSup_upper cSup_least)
+lemma cSup_upper2: "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> bdd_above X \<Longrightarrow> y \<le> Sup X"
+  by (metis cSup_upper order_trans)
+
+lemma cInf_lower2: "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> y"
+  by (metis cInf_lower order_trans)
+
+lemma cSup_mono: "B \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. b \<le> a) \<Longrightarrow> Sup B \<le> Sup A"
+  by (metis cSup_least cSup_upper2)
+
+lemma cInf_mono: "B \<noteq> {} \<Longrightarrow> bdd_below A \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<le> b) \<Longrightarrow> Inf A \<le> Inf B"
+  by (metis cInf_greatest cInf_lower2)
 
-lemma cInf_eq_minimum: (*REAL_INF_MIN in HOL4*)
-  "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X = z"
-  by (intro antisym cInf_lower[of z X z] cInf_greatest[of X z]) auto
+lemma cSup_subset_mono: "A \<noteq> {} \<Longrightarrow> bdd_above B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Sup A \<le> Sup B"
+  by (metis cSup_least cSup_upper subsetD)
+
+lemma cInf_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Inf B \<le> Inf A"
+  by (metis cInf_greatest cInf_lower subsetD)
 
-lemma cSup_le_iff: "S \<noteq> {} \<Longrightarrow> (\<And>a. a \<in> S \<Longrightarrow> a \<le> z) \<Longrightarrow> Sup S \<le> a \<longleftrightarrow> (\<forall>x\<in>S. x \<le> a)"
+lemma cSup_eq_maximum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
+  by (intro antisym cSup_upper[of z X] cSup_least[of X z]) auto
+
+lemma cInf_eq_minimum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X = z"
+  by (intro antisym cInf_lower[of z X] cInf_greatest[of X z]) auto
+
+lemma cSup_le_iff: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S \<le> a \<longleftrightarrow> (\<forall>x\<in>S. x \<le> a)"
   by (metis order_trans cSup_upper cSup_least)
 
-lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> (\<And>a. a \<in> S \<Longrightarrow> z \<le> a) \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
+lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
   by (metis order_trans cInf_lower cInf_greatest)
 
-lemma cSup_singleton [simp]: "Sup {x} = x"
-  by (intro cSup_eq_maximum) auto
-
-lemma cInf_singleton [simp]: "Inf {x} = x"
-  by (intro cInf_eq_minimum) auto
-
-lemma cSup_upper2: (*REAL_IMP_LE_SUP in HOL4*)
-  "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> y \<le> Sup X"
-  by (metis cSup_upper order_trans)
- 
-lemma cInf_lower2:
-  "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X \<le> y"
-  by (metis cInf_lower order_trans)
-
-lemma cSup_upper_EX: "x \<in> X \<Longrightarrow> \<exists>z. \<forall>x. x \<in> X \<longrightarrow> x \<le> z \<Longrightarrow> x \<le> Sup X"
-  by (blast intro: cSup_upper)
-
-lemma cInf_lower_EX:  "x \<in> X \<Longrightarrow> \<exists>z. \<forall>x. x \<in> X \<longrightarrow> z \<le> x \<Longrightarrow> Inf X \<le> x"
-  by (blast intro: cInf_lower)
-
 lemma cSup_eq_non_empty:
   assumes 1: "X \<noteq> {}"
   assumes 2: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
@@ -77,67 +214,47 @@
   shows "Inf X = a"
   by (intro 3 1 antisym cInf_greatest) (auto intro: 2 1 cInf_lower)
 
-lemma cInf_cSup: "S \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf S = Sup {x. \<forall>s\<in>S. x \<le> s}"
-  by (rule cInf_eq_non_empty) (auto intro: cSup_upper cSup_least)
+lemma cInf_cSup: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> Inf S = Sup {x. \<forall>s\<in>S. x \<le> s}"
+  by (rule cInf_eq_non_empty) (auto intro!: cSup_upper cSup_least simp: bdd_below_def)
 
-lemma cSup_cInf: "S \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup S = Inf {x. \<forall>s\<in>S. s \<le> x}"
-  by (rule cSup_eq_non_empty) (auto intro: cInf_lower cInf_greatest)
+lemma cSup_cInf: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S = Inf {x. \<forall>s\<in>S. s \<le> x}"
+  by (rule cSup_eq_non_empty) (auto intro!: cInf_lower cInf_greatest simp: bdd_above_def)
 
-lemma cSup_insert: 
-  assumes x: "X \<noteq> {}"
-      and z: "\<And>x. x \<in> X \<Longrightarrow> x \<le> z"
-  shows "Sup (insert a X) = sup a (Sup X)"
-proof (intro cSup_eq_non_empty)
-  fix y assume "\<And>x. x \<in> insert a X \<Longrightarrow> x \<le> y" with x show "sup a (Sup X) \<le> y" by (auto intro: cSup_least)
-qed (auto intro: le_supI2 z cSup_upper)
+lemma cSup_insert: "X \<noteq> {} \<Longrightarrow> bdd_above X \<Longrightarrow> Sup (insert a X) = sup a (Sup X)"
+  by (intro cSup_eq_non_empty) (auto intro: le_supI2 cSup_upper cSup_least)
+
+lemma cInf_insert: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf (insert a X) = inf a (Inf X)"
+  by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest)
 
-lemma cInf_insert: 
-  assumes x: "X \<noteq> {}"
-      and z: "\<And>x. x \<in> X \<Longrightarrow> z \<le> x"
-  shows "Inf (insert a X) = inf a (Inf X)"
-proof (intro cInf_eq_non_empty)
-  fix y assume "\<And>x. x \<in> insert a X \<Longrightarrow> y \<le> x" with x show "y \<le> inf a (Inf X)" by (auto intro: cInf_greatest)
-qed (auto intro: le_infI2 z cInf_lower)
+lemma cSup_singleton [simp]: "Sup {x} = x"
+  by (intro cSup_eq_maximum) auto
+
+lemma cInf_singleton [simp]: "Inf {x} = x"
+  by (intro cInf_eq_minimum) auto
 
-lemma cSup_insert_If: 
-  "(\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup (insert a X) = (if X = {} then a else sup a (Sup X))"
-  using cSup_insert[of X z] by simp
+lemma cSup_insert_If:  "bdd_above X \<Longrightarrow> Sup (insert a X) = (if X = {} then a else sup a (Sup X))"
+  using cSup_insert[of X] by simp
 
-lemma cInf_insert_if: 
-  "(\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
-  using cInf_insert[of X z] by simp
+lemma cInf_insert_If: "bdd_below X \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
+  using cInf_insert[of X] by simp
 
 lemma le_cSup_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> x \<le> Sup X"
 proof (induct X arbitrary: x rule: finite_induct)
   case (insert x X y) then show ?case
-    apply (cases "X = {}")
-    apply simp
-    apply (subst cSup_insert[of _ "Sup X"])
-    apply (auto intro: le_supI2)
-    done
+    by (cases "X = {}") (auto simp: cSup_insert intro: le_supI2)
 qed simp
 
 lemma cInf_le_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> Inf X \<le> x"
 proof (induct X arbitrary: x rule: finite_induct)
   case (insert x X y) then show ?case
-    apply (cases "X = {}")
-    apply simp
-    apply (subst cInf_insert[of _ "Inf X"])
-    apply (auto intro: le_infI2)
-    done
+    by (cases "X = {}") (auto simp: cInf_insert intro: le_infI2)
 qed simp
 
 lemma cSup_eq_Sup_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Sup_fin X"
-proof (induct X rule: finite_ne_induct)
-  case (insert x X) then show ?case
-    using cSup_insert[of X "Sup_fin X" x] le_cSup_finite[of X] by simp
-qed simp
+  by (induct X rule: finite_ne_induct) (simp_all add: cSup_insert)
 
 lemma cInf_eq_Inf_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Inf_fin X"
-proof (induct X rule: finite_ne_induct)
-  case (insert x X) then show ?case
-    using cInf_insert[of X "Inf_fin X" x] cInf_le_finite[of X] by simp
-qed simp
+  by (induct X rule: finite_ne_induct) (simp_all add: cInf_insert)
 
 lemma cSup_atMost[simp]: "Sup {..x} = x"
   by (auto intro!: cSup_eq_maximum)
@@ -157,16 +274,91 @@
 lemma cInf_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Inf {y..x} = y"
   by (auto intro!: cInf_eq_minimum)
 
+lemma cINF_lower: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> INFI A f \<le> f x"
+  unfolding INF_def by (rule cInf_lower) auto
+
+lemma cINF_greatest: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> m \<le> f x) \<Longrightarrow> m \<le> INFI A f"
+  unfolding INF_def by (rule cInf_greatest) auto
+
+lemma cSUP_upper: "x \<in> A \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> f x \<le> SUPR A f"
+  unfolding SUP_def by (rule cSup_upper) auto
+
+lemma cSUP_least: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<le> M) \<Longrightarrow> SUPR A f \<le> M"
+  unfolding SUP_def by (rule cSup_least) auto
+
+lemma cINF_lower2: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> f x \<le> u \<Longrightarrow> INFI A f \<le> u"
+  by (auto intro: cINF_lower assms order_trans)
+
+lemma cSUP_upper2: "bdd_above (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> u \<le> f x \<Longrightarrow> u \<le> SUPR A f"
+  by (auto intro: cSUP_upper assms order_trans)
+
+lemma cSUP_const: "A \<noteq> {} \<Longrightarrow> (SUP x:A. c) = c"
+  by (intro antisym cSUP_least) (auto intro: cSUP_upper)
+
+lemma cINF_const: "A \<noteq> {} \<Longrightarrow> (INF x:A. c) = c"
+  by (intro antisym cINF_greatest) (auto intro: cINF_lower)
+
+lemma le_cINF_iff: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> u \<le> INFI A f \<longleftrightarrow> (\<forall>x\<in>A. u \<le> f x)"
+  by (metis cINF_greatest cINF_lower assms order_trans)
+
+lemma cSUP_le_iff: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> SUPR A f \<le> u \<longleftrightarrow> (\<forall>x\<in>A. f x \<le> u)"
+  by (metis cSUP_least cSUP_upper assms order_trans)
+
+lemma less_cINF_D: "bdd_below (f`A) \<Longrightarrow> y < (INF i:A. f i) \<Longrightarrow> i \<in> A \<Longrightarrow> y < f i"
+  by (metis cINF_lower less_le_trans)
+
+lemma cSUP_lessD: "bdd_above (f`A) \<Longrightarrow> (SUP i:A. f i) < y \<Longrightarrow> i \<in> A \<Longrightarrow> f i < y"
+  by (metis cSUP_upper le_less_trans)
+
+lemma cINF_insert: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> INFI (insert a A) f = inf (f a) (INFI A f)"
+  by (metis INF_def cInf_insert assms empty_is_image image_insert)
+
+lemma cSUP_insert: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> SUPR (insert a A) f = sup (f a) (SUPR A f)"
+  by (metis SUP_def cSup_insert assms empty_is_image image_insert)
+
+lemma cINF_mono: "B \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> (\<And>m. m \<in> B \<Longrightarrow> \<exists>n\<in>A. f n \<le> g m) \<Longrightarrow> INFI A f \<le> INFI B g"
+  unfolding INF_def by (auto intro: cInf_mono)
+
+lemma cSUP_mono: "A \<noteq> {} \<Longrightarrow> bdd_above (g ` B) \<Longrightarrow> (\<And>n. n \<in> A \<Longrightarrow> \<exists>m\<in>B. f n \<le> g m) \<Longrightarrow> SUPR A f \<le> SUPR B g"
+  unfolding SUP_def by (auto intro: cSup_mono)
+
+lemma cINF_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below (g ` B) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> g x \<le> f x) \<Longrightarrow> INFI B g \<le> INFI A f"
+  by (rule cINF_mono) auto
+
+lemma cSUP_subset_mono: "A \<noteq> {} \<Longrightarrow> bdd_above (g ` B) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> f x \<le> g x) \<Longrightarrow> SUPR A f \<le> SUPR B g"
+  by (rule cSUP_mono) auto
+
+lemma less_eq_cInf_inter: "bdd_below A \<Longrightarrow> bdd_below B \<Longrightarrow> A \<inter> B \<noteq> {} \<Longrightarrow> inf (Inf A) (Inf B) \<le> Inf (A \<inter> B)"
+  by (metis cInf_superset_mono lattice_class.inf_sup_ord(1) le_infI1)
+
+lemma cSup_inter_less_eq: "bdd_above A \<Longrightarrow> bdd_above B \<Longrightarrow> A \<inter> B \<noteq> {} \<Longrightarrow> Sup (A \<inter> B) \<le> sup (Sup A) (Sup B) "
+  by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1)
+
+lemma cInf_union_distrib: "A \<noteq> {} \<Longrightarrow> bdd_below A \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> Inf (A \<union> B) = inf (Inf A) (Inf B)"
+  by (intro antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower)
+
+lemma cINF_union: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_below (f`B) \<Longrightarrow> INFI (A \<union> B) f = inf (INFI A f) (INFI B f)"
+  unfolding INF_def using assms by (auto simp add: image_Un intro: cInf_union_distrib)
+
+lemma cSup_union_distrib: "A \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_above B \<Longrightarrow> Sup (A \<union> B) = sup (Sup A) (Sup B)"
+  by (intro antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper)
+
+lemma cSUP_union: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_above (f`B) \<Longrightarrow> SUPR (A \<union> B) f = sup (SUPR A f) (SUPR B f)"
+  unfolding SUP_def by (auto simp add: image_Un intro: cSup_union_distrib)
+
+lemma cINF_inf_distrib: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> bdd_below (g`A) \<Longrightarrow> inf (INFI A f) (INFI A g) = (INF a:A. inf (f a) (g a))"
+  by (intro antisym le_infI cINF_greatest cINF_lower2)
+     (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI)
+
+lemma SUP_sup_distrib: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> bdd_above (g`A) \<Longrightarrow> sup (SUPR A f) (SUPR A g) = (SUP a:A. sup (f a) (g a))"
+  by (intro antisym le_supI cSUP_least cSUP_upper2)
+     (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI)
+
 end
 
 instance complete_lattice \<subseteq> conditionally_complete_lattice
   by default (auto intro: Sup_upper Sup_least Inf_lower Inf_greatest)
 
-lemma isLub_cSup: 
-  "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> (\<exists>b. S *<= b) \<Longrightarrow> isLub UNIV S (Sup S)"
-  by  (auto simp add: isLub_def setle_def leastP_def isUb_def
-            intro!: setgeI intro: cSup_upper cSup_least)
-
 lemma cSup_eq:
   fixes a :: "'a :: {conditionally_complete_lattice, no_bot}"
   assumes upper: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
@@ -185,33 +377,33 @@
   assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
 qed (intro cInf_eq_non_empty assms)
 
-lemma cSup_le: "(S::'a::conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> S *<= b \<Longrightarrow> Sup S \<le> b"
-  by (metis cSup_least setle_def)
-
-lemma cInf_ge: "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> b <=* S \<Longrightarrow> Inf S \<ge> b"
-  by (metis cInf_greatest setge_def)
-
 class conditionally_complete_linorder = conditionally_complete_lattice + linorder
 begin
 
 lemma less_cSup_iff : (*REAL_SUP_LE in HOL4*)
-  "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> y < Sup X \<longleftrightarrow> (\<exists>x\<in>X. y < x)"
+  "X \<noteq> {} \<Longrightarrow> bdd_above X \<Longrightarrow> y < Sup X \<longleftrightarrow> (\<exists>x\<in>X. y < x)"
   by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans)
 
-lemma cInf_less_iff: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X < y \<longleftrightarrow> (\<exists>x\<in>X. x < y)"
+lemma cInf_less_iff: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X < y \<longleftrightarrow> (\<exists>x\<in>X. x < y)"
   by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans)
 
+lemma cINF_less_iff: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> (INF i:A. f i) < a \<longleftrightarrow> (\<exists>x\<in>A. f x < a)"
+  unfolding INF_def using cInf_less_iff[of "f`A"] by auto
+
+lemma less_cSUP_iff: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> a < (SUP i:A. f i) \<longleftrightarrow> (\<exists>x\<in>A. a < f x)"
+  unfolding SUP_def using less_cSup_iff[of "f`A"] by auto
+
 lemma less_cSupE:
   assumes "y < Sup X" "X \<noteq> {}" obtains x where "x \<in> X" "y < x"
   by (metis cSup_least assms not_le that)
 
 lemma less_cSupD:
   "X \<noteq> {} \<Longrightarrow> z < Sup X \<Longrightarrow> \<exists>x\<in>X. z < x"
-  by (metis less_cSup_iff not_leE)
+  by (metis less_cSup_iff not_leE bdd_above_def)
 
 lemma cInf_lessD:
   "X \<noteq> {} \<Longrightarrow> Inf X < z \<Longrightarrow> \<exists>x\<in>X. x < z"
-  by (metis cInf_less_iff not_leE)
+  by (metis cInf_less_iff not_leE bdd_below_def)
 
 lemma complete_interval:
   assumes "a < b" and "P a" and "\<not> P b"
@@ -219,7 +411,7 @@
              (\<forall>d. (\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x) \<longrightarrow> d \<le> c)"
 proof (rule exI [where x = "Sup {d. \<forall>x. a \<le> x & x < d --> P x}"], auto)
   show "a \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
-    by (rule cSup_upper [where z=b], auto)
+    by (rule cSup_upper, auto simp: bdd_above_def)
        (metis `a < b` `\<not> P b` linear less_le)
 next
   show "Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c} \<le> b"
@@ -240,12 +432,36 @@
   fix d
     assume 0: "\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x"
     thus "d \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
-      by (rule_tac z="b" in cSup_upper, auto) 
+      by (rule_tac cSup_upper, auto simp: bdd_above_def)
          (metis `a<b` `~ P b` linear less_le)
 qed
 
 end
 
+lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
+  using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp
+
+lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
+  using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp
+
+lemma cSup_lessThan[simp]: "Sup {..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
+  by (auto intro!: cSup_eq_non_empty intro: dense_le)
+
+lemma cSup_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Sup {y<..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
+  by (auto intro!: cSup_eq intro: dense_le_bounded)
+
+lemma cSup_atLeastLessThan[simp]: "y < x \<Longrightarrow> Sup {y..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
+  by (auto intro!: cSup_eq intro: dense_le_bounded)
+
+lemma cInf_greaterThan[simp]: "Inf {x::'a::{conditionally_complete_linorder, no_top, dense_linorder} <..} = x"
+  by (auto intro!: cInf_eq intro: dense_ge)
+
+lemma cInf_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Inf {y<..x::'a::{conditionally_complete_linorder, no_top, dense_linorder}} = y"
+  by (auto intro!: cInf_eq intro: dense_ge_bounded)
+
+lemma cInf_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Inf {y<..<x::'a::{conditionally_complete_linorder, no_top, dense_linorder}} = y"
+  by (auto intro!: cInf_eq intro: dense_ge_bounded)
+
 class linear_continuum = conditionally_complete_linorder + dense_linorder +
   assumes UNIV_not_singleton: "\<exists>a b::'a. a \<noteq> b"
 begin
@@ -255,50 +471,92 @@
 
 end
 
-lemma cSup_bounds:
-  fixes S :: "'a :: conditionally_complete_lattice set"
-  assumes Se: "S \<noteq> {}" and l: "a <=* S" and u: "S *<= b"
-  shows "a \<le> Sup S \<and> Sup S \<le> b"
-proof-
-  from isLub_cSup[OF Se] u have lub: "isLub UNIV S (Sup S)" by blast
-  hence b: "Sup S \<le> b" using u 
-    by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def) 
-  from Se obtain y where y: "y \<in> S" by blast
-  from lub l have "a \<le> Sup S"
-    by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def)
-       (metis le_iff_sup le_sup_iff y)
-  with b show ?thesis by blast
+instantiation nat :: conditionally_complete_linorder
+begin
+
+definition "Sup (X::nat set) = Max X"
+definition "Inf (X::nat set) = (LEAST n. n \<in> X)"
+
+lemma bdd_above_nat: "bdd_above X \<longleftrightarrow> finite (X::nat set)"
+proof
+  assume "bdd_above X"
+  then obtain z where "X \<subseteq> {.. z}"
+    by (auto simp: bdd_above_def)
+  then show "finite X"
+    by (rule finite_subset) simp
+qed simp
+
+instance
+proof
+  fix x :: nat and X :: "nat set"
+  { assume "x \<in> X" "bdd_below X" then show "Inf X \<le> x"
+      by (simp add: Inf_nat_def Least_le) }
+  { assume "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> x \<le> y" then show "x \<le> Inf X"
+      unfolding Inf_nat_def ex_in_conv[symmetric] by (rule LeastI2_ex) }
+  { assume "x \<in> X" "bdd_above X" then show "x \<le> Sup X"
+      by (simp add: Sup_nat_def bdd_above_nat) }
+  { assume "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> y \<le> x" 
+    moreover then have "bdd_above X"
+      by (auto simp: bdd_above_def)
+    ultimately show "Sup X \<le> x"
+      by (simp add: Sup_nat_def bdd_above_nat) }
 qed
-
+end
 
-lemma cSup_unique: "(S::'a :: {conditionally_complete_linorder, no_bot} set) *<= b \<Longrightarrow> (\<forall>b'<b. \<exists>x\<in>S. b' < x) \<Longrightarrow> Sup S = b"
-  by (rule cSup_eq) (auto simp: not_le[symmetric] setle_def)
+instantiation int :: conditionally_complete_linorder
+begin
 
-lemma cInf_unique: "b <=* (S::'a :: {conditionally_complete_linorder, no_top} set) \<Longrightarrow> (\<forall>b'>b. \<exists>x\<in>S. b' > x) \<Longrightarrow> Inf S = b"
-  by (rule cInf_eq) (auto simp: not_le[symmetric] setge_def)
+definition "Sup (X::int set) = (THE x. x \<in> X \<and> (\<forall>y\<in>X. y \<le> x))"
+definition "Inf (X::int set) = - (Sup (uminus ` X))"
 
-lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
-  using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp
-
-lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
-  using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp
-
-lemma cSup_lessThan[simp]: "Sup {..<x::'a::{conditionally_complete_linorder, unbounded_dense_linorder}} = x"
-  by (auto intro!: cSup_eq_non_empty intro: dense_le)
-
-lemma cSup_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Sup {y<..<x::'a::{conditionally_complete_linorder, unbounded_dense_linorder}} = x"
-  by (auto intro!: cSup_eq intro: dense_le_bounded)
+instance
+proof
+  { fix x :: int and X :: "int set" assume "X \<noteq> {}" "bdd_above X"
+    then obtain x y where "X \<subseteq> {..y}" "x \<in> X"
+      by (auto simp: bdd_above_def)
+    then have *: "finite (X \<inter> {x..y})" "X \<inter> {x..y} \<noteq> {}" and "x \<le> y"
+      by (auto simp: subset_eq)
+    have "\<exists>!x\<in>X. (\<forall>y\<in>X. y \<le> x)"
+    proof
+      { fix z assume "z \<in> X"
+        have "z \<le> Max (X \<inter> {x..y})"
+        proof cases
+          assume "x \<le> z" with `z \<in> X` `X \<subseteq> {..y}` *(1) show ?thesis
+            by (auto intro!: Max_ge)
+        next
+          assume "\<not> x \<le> z"
+          then have "z < x" by simp
+          also have "x \<le> Max (X \<inter> {x..y})"
+            using `x \<in> X` *(1) `x \<le> y` by (intro Max_ge) auto
+          finally show ?thesis by simp
+        qed }
+      note le = this
+      with Max_in[OF *] show ex: "Max (X \<inter> {x..y}) \<in> X \<and> (\<forall>z\<in>X. z \<le> Max (X \<inter> {x..y}))" by auto
 
-lemma cSup_atLeastLessThan[simp]: "y < x \<Longrightarrow> Sup {y..<x::'a::{conditionally_complete_linorder, unbounded_dense_linorder}} = x"
-  by (auto intro!: cSup_eq intro: dense_le_bounded)
-
-lemma cInf_greaterThan[simp]: "Inf {x::'a::{conditionally_complete_linorder, unbounded_dense_linorder} <..} = x"
-  by (auto intro!: cInf_eq intro: dense_ge)
+      fix z assume *: "z \<in> X \<and> (\<forall>y\<in>X. y \<le> z)"
+      with le have "z \<le> Max (X \<inter> {x..y})"
+        by auto
+      moreover have "Max (X \<inter> {x..y}) \<le> z"
+        using * ex by auto
+      ultimately show "z = Max (X \<inter> {x..y})"
+        by auto
+    qed
+    then have "Sup X \<in> X \<and> (\<forall>y\<in>X. y \<le> Sup X)"
+      unfolding Sup_int_def by (rule theI') }
+  note Sup_int = this
 
-lemma cInf_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Inf {y<..x::'a::{conditionally_complete_linorder, unbounded_dense_linorder}} = y"
-  by (auto intro!: cInf_eq intro: dense_ge_bounded)
+  { fix x :: int and X :: "int set" assume "x \<in> X" "bdd_above X" then show "x \<le> Sup X"
+      using Sup_int[of X] by auto }
+  note le_Sup = this
+  { fix x :: int and X :: "int set" assume "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> y \<le> x" then show "Sup X \<le> x"
+      using Sup_int[of X] by (auto simp: bdd_above_def) }
+  note Sup_le = this
 
-lemma cInf_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Inf {y<..<x::'a::{conditionally_complete_linorder, unbounded_dense_linorder}} = y"
-  by (auto intro!: cInf_eq intro: dense_ge_bounded)
+  { fix x :: int and X :: "int set" assume "x \<in> X" "bdd_below X" then show "Inf X \<le> x"
+      using le_Sup[of "-x" "uminus ` X"] by (auto simp: Inf_int_def) }
+  { fix x :: int and X :: "int set" assume "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> x \<le> y" then show "x \<le> Inf X"
+      using Sup_le[of "uminus ` X" "-x"] by (force simp: Inf_int_def) }
+qed
+end
 
 end
--- a/src/HOL/Decision_Procs/Approximation.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Decision_Procs/Approximation.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -29,7 +29,7 @@
   have shift_pow: "\<And>i. - (x * ((-1)^i * a (Suc i) * x ^ i)) = (-1)^(Suc i) * a (Suc i) * x ^ (Suc i)"
     by auto
   show ?thesis
-    unfolding setsum_right_distrib shift_pow diff_minus setsum_negf[symmetric]
+    unfolding setsum_right_distrib shift_pow uminus_add_conv_diff [symmetric] setsum_negf[symmetric]
     setsum_head_upt_Suc[OF zero_less_Suc]
     setsum_reindex[OF inj_Suc, unfolded comp_def, symmetric, of "\<lambda> n. (-1)^n  *a n * x^n"] by auto
 qed
@@ -132,14 +132,7 @@
 lemma get_odd[simp]: "odd (get_odd n)" unfolding get_odd_def by (cases "odd n", auto)
 lemma get_even[simp]: "even (get_even n)" unfolding get_even_def by (cases "even n", auto)
 lemma get_odd_ex: "\<exists> k. Suc k = get_odd n \<and> odd (Suc k)"
-proof (cases "odd n")
-  case True hence "0 < n" by (rule odd_pos)
-  from gr0_implies_Suc[OF this] obtain k where "Suc k = n" by auto
-  thus ?thesis unfolding get_odd_def if_P[OF True] using True[unfolded `Suc k = n`[symmetric]] by blast
-next
-  case False hence "odd (Suc n)" by auto
-  thus ?thesis unfolding get_odd_def if_not_P[OF False] by blast
-qed
+  by (auto simp: get_odd_def odd_pos intro!: exI[of _ "n - 1"])
 
 lemma get_even_double: "\<exists>i. get_even n = 2 * i" using get_even[unfolded even_mult_two_ex] .
 lemma get_odd_double: "\<exists>i. get_odd n = 2 * i + 1" using get_odd[unfolded odd_Suc_mult_two_ex] by auto
@@ -151,47 +144,9 @@
                       else if u < 0         then (u ^ n, l ^ n)
                                             else (0, (max (-l) u) ^ n))"
 
-lemma float_power_bnds: fixes x :: real
-  assumes "(l1, u1) = float_power_bnds n l u" and "x \<in> {l .. u}"
-  shows "x ^ n \<in> {l1..u1}"
-proof (cases "even n")
-  case True
-  show ?thesis
-  proof (cases "0 < l")
-    case True hence "odd n \<or> 0 < l" and "0 \<le> real l" by auto
-    have u1: "u1 = u ^ n" and l1: "l1 = l ^ n" using assms
-      unfolding float_power_bnds_def if_P[OF `odd n \<or> 0 < l`] by auto
-    have "real l ^ n \<le> x ^ n" and "x ^ n \<le> real u ^ n " using `0 \<le> real l` assms
-      by (auto simp: power_mono)
-    thus ?thesis using assms `0 < l` unfolding l1 u1 by auto
-  next
-    case False hence P: "\<not> (odd n \<or> 0 < l)" using `even n` by auto
-    show ?thesis
-    proof (cases "u < 0")
-      case True hence "0 \<le> - real u" and "- real u \<le> - x" and "0 \<le> - x" and "-x \<le> - real l" using assms  by auto
-      hence "real u ^ n \<le> x ^ n" and "x ^ n \<le> real l ^ n" using power_mono[of  "-x" "-real l" n] power_mono[of "-real u" "-x" n]
-        unfolding power_minus_even[OF `even n`] by auto
-      moreover have u1: "u1 = l ^ n" and l1: "l1 = u ^ n" using assms unfolding float_power_bnds_def if_not_P[OF P] if_P[OF True] by auto
-      ultimately show ?thesis by auto
-    next
-      case False
-      have "\<bar>x\<bar> \<le> real (max (-l) u)"
-      proof (cases "-l \<le> u")
-        case True thus ?thesis unfolding max_def if_P[OF True] using assms by auto
-      next
-        case False thus ?thesis unfolding max_def if_not_P[OF False] using assms by auto
-      qed
-      hence x_abs: "\<bar>x\<bar> \<le> \<bar>real (max (-l) u)\<bar>" by auto
-      have u1: "u1 = (max (-l) u) ^ n" and l1: "l1 = 0" using assms unfolding float_power_bnds_def if_not_P[OF P] if_not_P[OF False] by auto
-      show ?thesis unfolding atLeastAtMost_iff l1 u1 using zero_le_even_power[OF `even n`] power_mono_even[OF `even n` x_abs] by auto
-    qed
-  qed
-next
-  case False hence "odd n \<or> 0 < l" by auto
-  have u1: "u1 = u ^ n" and l1: "l1 = l ^ n" using assms unfolding float_power_bnds_def if_P[OF `odd n \<or> 0 < l`] by auto
-  have "real l ^ n \<le> x ^ n" and "x ^ n \<le> real u ^ n " using assms unfolding atLeastAtMost_iff using power_mono_odd[OF False] by auto
-  thus ?thesis unfolding atLeastAtMost_iff l1 u1 less_float_def by auto
-qed
+lemma float_power_bnds: "(l1, u1) = float_power_bnds n l u \<Longrightarrow> x \<in> {l .. u} \<Longrightarrow> (x::real) ^ n \<in> {l1..u1}"
+  by (auto simp: float_power_bnds_def max_def split: split_if_asm
+           intro: power_mono_odd power_mono power_mono_even zero_le_even_power)
 
 lemma bnds_power: "\<forall> (x::real) l u. (l1, u1) = float_power_bnds n l u \<and> x \<in> {l .. u} \<longrightarrow> l1 \<le> x ^ n \<and> x ^ n \<le> u1"
   using float_power_bnds by auto
@@ -244,7 +199,7 @@
 qed
 
 lemma sqrt_iteration_bound: assumes "0 < real x"
-  shows "sqrt x < (sqrt_iteration prec n x)"
+  shows "sqrt x < sqrt_iteration prec n x"
 proof (induct n)
   case 0
   show ?case
@@ -356,20 +311,8 @@
   note ub = this
 
   show ?thesis
-  proof (cases "0 < x")
-    case True with lb ub show ?thesis by auto
-  next case False show ?thesis
-  proof (cases "real x = 0")
-    case True thus ?thesis
-      by (auto simp add: lb_sqrt.simps ub_sqrt.simps)
-  next
-    case False with `\<not> 0 < x` have "x < 0" and "0 < -x"
-      by auto
-
-    with `\<not> 0 < x`
-    show ?thesis using lb[OF `0 < -x`] ub[OF `0 < -x`]
-      by (auto simp add: real_sqrt_minus lb_sqrt.simps ub_sqrt.simps)
-  qed qed
+    using lb[of "-x"] ub[of "-x"] lb[of x] ub[of x]
+    by (auto simp add: lb_sqrt.simps ub_sqrt.simps real_sqrt_minus)
 qed
 
 lemma bnds_sqrt: "\<forall> (x::real) lx ux. (l, u) = (lb_sqrt prec lx, ub_sqrt prec ux) \<and> x \<in> {lx .. ux} \<longrightarrow> l \<le> sqrt x \<and> sqrt x \<le> u"
@@ -412,8 +355,8 @@
   assumes "0 \<le> real x" "real x \<le> 1" and "even n"
   shows "arctan x \<in> {(x * lb_arctan_horner prec n 1 (x * x)) .. (x * ub_arctan_horner prec (Suc n) 1 (x * x))}"
 proof -
-  let "?c i" = "-1^i * (1 / (i * 2 + (1::nat)) * real x ^ (i * 2 + 1))"
-  let "?S n" = "\<Sum> i=0..<n. ?c i"
+  let ?c = "\<lambda>i. -1^i * (1 / (i * 2 + (1::nat)) * real x ^ (i * 2 + 1))"
+  let ?S = "\<lambda>n. \<Sum> i=0..<n. ?c i"
 
   have "0 \<le> real (x * x)" by auto
   from `even n` obtain m where "2 * m = n" unfolding even_mult_two_ex by auto
@@ -457,30 +400,11 @@
 
 lemma arctan_0_1_bounds: assumes "0 \<le> real x" "real x \<le> 1"
   shows "arctan x \<in> {(x * lb_arctan_horner prec (get_even n) 1 (x * x)) .. (x * ub_arctan_horner prec (get_odd n) 1 (x * x))}"
-proof (cases "even n")
-  case True
-  obtain n' where "Suc n' = get_odd n" and "odd (Suc n')" using get_odd_ex by auto
-  hence "even n'" unfolding even_Suc by auto
-  have "arctan x \<le> x * ub_arctan_horner prec (get_odd n) 1 (x * x)"
-    unfolding `Suc n' = get_odd n`[symmetric] using arctan_0_1_bounds'[OF `0 \<le> real x` `real x \<le> 1` `even n'`] by auto
-  moreover
-  have "x * lb_arctan_horner prec (get_even n) 1 (x * x) \<le> arctan x"
-    unfolding get_even_def if_P[OF True] using arctan_0_1_bounds'[OF `0 \<le> real x` `real x \<le> 1` `even n`] by auto
-  ultimately show ?thesis by auto
-next
-  case False hence "0 < n" by (rule odd_pos)
-  from gr0_implies_Suc[OF this] obtain n' where "n = Suc n'" ..
-  from False[unfolded this even_Suc]
-  have "even n'" and "even (Suc (Suc n'))" by auto
-  have "get_odd n = Suc n'" unfolding get_odd_def if_P[OF False] using `n = Suc n'` .
-
-  have "arctan x \<le> x * ub_arctan_horner prec (get_odd n) 1 (x * x)"
-    unfolding `get_odd n = Suc n'` using arctan_0_1_bounds'[OF `0 \<le> real x` `real x \<le> 1` `even n'`] by auto
-  moreover
-  have "(x * lb_arctan_horner prec (get_even n) 1 (x * x)) \<le> arctan x"
-    unfolding get_even_def if_not_P[OF False] unfolding `n = Suc n'` using arctan_0_1_bounds'[OF `0 \<le> real x` `real x \<le> 1` `even (Suc (Suc n'))`] by auto
-  ultimately show ?thesis by auto
-qed
+  using
+    arctan_0_1_bounds'[OF assms, of n prec]
+    arctan_0_1_bounds'[OF assms, of "n + 1" prec]
+    arctan_0_1_bounds'[OF assms, of "n - 1" prec]
+  by (auto simp: get_even_def get_odd_def odd_pos simp del: ub_arctan_horner.simps lb_arctan_horner.simps)
 
 subsection "Compute \<pi>"
 
@@ -530,16 +454,11 @@
     finally have "?k * lb_arctan_horner prec (get_even n) 1 (?k * ?k) \<le> arctan (1 / k)" .
   } note lb_arctan = this
 
-  have "pi \<le> ub_pi n"
-    unfolding ub_pi_def machin_pi Let_def unfolding Float_num
-    using lb_arctan[of 239] ub_arctan[of 5] powr_realpow[of 2 2]
-    by (auto intro!: mult_left_mono add_mono simp add: diff_minus)
-  moreover
-  have "lb_pi n \<le> pi"
-    unfolding lb_pi_def machin_pi Let_def Float_num
-    using lb_arctan[of 5] ub_arctan[of 239] powr_realpow[of 2 2]
-    by (auto intro!: mult_left_mono add_mono simp add: diff_minus)
-  ultimately show ?thesis by auto
+  have "pi \<le> ub_pi n \<and> lb_pi n \<le> pi"
+    unfolding lb_pi_def ub_pi_def machin_pi Let_def unfolding Float_num
+    using lb_arctan[of 5] ub_arctan[of 239] lb_arctan[of 239] ub_arctan[of 5] powr_realpow[of 2 2]
+    by (auto intro!: mult_left_mono add_mono simp add: uminus_add_conv_diff [symmetric] simp del: uminus_add_conv_diff)
+  then show ?thesis by auto
 qed
 
 subsection "Compute arcus tangens in the entire domain"
@@ -1208,8 +1127,8 @@
     using x unfolding k[symmetric]
     by (cases "k = 0")
        (auto intro!: add_mono
-                simp add: diff_minus k[symmetric]
-                simp del: float_of_numeral)
+                simp add: k [symmetric] uminus_add_conv_diff [symmetric]
+                simp del: float_of_numeral uminus_add_conv_diff)
   note lx = this[THEN conjunct1] and ux = this[THEN conjunct2]
   hence lx_less_ux: "?lx \<le> real ?ux" by (rule order_trans)
 
@@ -1223,7 +1142,7 @@
     also have "\<dots> \<le> cos (x + (-k) * (2 * pi))"
       using cos_monotone_minus_pi_0'[OF pi_lx lx x_le_0]
       by (simp only: uminus_float.rep_eq real_of_int_minus
-        cos_minus diff_minus mult_minus_left)
+        cos_minus mult_minus_left) simp
     finally have "(lb_cos prec (- ?lx)) \<le> cos x"
       unfolding cos_periodic_int . }
   note negative_lx = this
@@ -1236,7 +1155,7 @@
     have "cos (x + (-k) * (2 * pi)) \<le> cos ?lx"
       using cos_monotone_0_pi'[OF lx_0 lx pi_x]
       by (simp only: real_of_int_minus
-        cos_minus diff_minus mult_minus_left)
+        cos_minus mult_minus_left) simp
     also have "\<dots> \<le> (ub_cos prec ?lx)"
       using lb_cos[OF lx_0 pi_lx] by simp
     finally have "cos x \<le> (ub_cos prec ?lx)"
@@ -1251,7 +1170,7 @@
     have "cos (x + (-k) * (2 * pi)) \<le> cos (real (- ?ux))"
       using cos_monotone_minus_pi_0'[OF pi_x ux ux_0]
       by (simp only: uminus_float.rep_eq real_of_int_minus
-          cos_minus diff_minus mult_minus_left)
+          cos_minus mult_minus_left) simp
     also have "\<dots> \<le> (ub_cos prec (- ?ux))"
       using lb_cos_minus[OF pi_ux ux_0, of prec] by simp
     finally have "cos x \<le> (ub_cos prec (- ?ux))"
@@ -1268,7 +1187,7 @@
     also have "\<dots> \<le> cos (x + (-k) * (2 * pi))"
       using cos_monotone_0_pi'[OF x_ge_0 ux pi_ux]
       by (simp only: real_of_int_minus
-        cos_minus diff_minus mult_minus_left)
+        cos_minus mult_minus_left) simp
     finally have "(lb_cos prec ?ux) \<le> cos x"
       unfolding cos_periodic_int . }
   note positive_ux = this
@@ -1343,7 +1262,7 @@
       also have "\<dots> \<le> cos ((?ux - 2 * ?lpi))"
         using cos_monotone_minus_pi_0'[OF pi_x x_le_ux ux_0]
         by (simp only: minus_float.rep_eq real_of_int_minus real_of_one minus_one[symmetric]
-            diff_minus mult_minus_left mult_1_left)
+            mult_minus_left mult_1_left) simp
       also have "\<dots> = cos ((- (?ux - 2 * ?lpi)))"
         unfolding uminus_float.rep_eq cos_minus ..
       also have "\<dots> \<le> (ub_cos prec (- (?ux - 2 * ?lpi)))"
@@ -1387,7 +1306,7 @@
       also have "\<dots> \<le> cos ((?lx + 2 * ?lpi))"
         using cos_monotone_0_pi'[OF lx_0 lx_le_x pi_x]
         by (simp only: minus_float.rep_eq real_of_int_minus real_of_one
-          minus_one[symmetric] diff_minus mult_minus_left mult_1_left)
+          minus_one[symmetric] mult_minus_left mult_1_left) simp
       also have "\<dots> \<le> (ub_cos prec (?lx + 2 * ?lpi))"
         using lb_cos[OF lx_0 pi_lx] by simp
       finally show ?thesis unfolding u by (simp add: real_of_float_max)
@@ -1808,10 +1727,8 @@
   done
 
 lemma Float_pos_eq_mantissa_pos:  "Float m e > 0 \<longleftrightarrow> m > 0"
-  apply (auto simp add: zero_less_mult_iff zero_float_def powr_gt_zero[of 2 "exponent x"] dest: less_zeroE)
   using powr_gt_zero[of 2 "e"]
-  apply simp
-  done
+  by (auto simp add: zero_less_mult_iff zero_float_def simp del: powr_gt_zero dest: less_zeroE)
 
 lemma Float_representation_aux:
   fixes m e
@@ -2164,12 +2081,12 @@
   unfolding divide_inverse interpret_floatarith.simps ..
 
 lemma interpret_floatarith_diff: "interpret_floatarith (Add a (Minus b)) vs = (interpret_floatarith a vs) - (interpret_floatarith b vs)"
-  unfolding diff_minus interpret_floatarith.simps ..
+  unfolding interpret_floatarith.simps by simp
 
 lemma interpret_floatarith_sin: "interpret_floatarith (Cos (Add (Mult Pi (Num (Float 1 -1))) (Minus a))) vs =
   sin (interpret_floatarith a vs)"
   unfolding sin_cos_eq interpret_floatarith.simps
-            interpret_floatarith_divide interpret_floatarith_diff diff_minus
+            interpret_floatarith_divide interpret_floatarith_diff
   by auto
 
 lemma interpret_floatarith_tan:
@@ -3192,7 +3109,7 @@
 
   from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x'
   have "ly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"
-    by (auto simp add: diff_minus)
+    by auto
   from order_less_le_trans[OF _ this, of 0] `0 < ly`
   show ?thesis by auto
 qed
@@ -3214,7 +3131,7 @@
 
   from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x'
   have "ly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"
-    by (auto simp add: diff_minus)
+    by auto
   from order_trans[OF _ this, of 0] `0 \<le> ly`
   show ?thesis by auto
 qed
--- a/src/HOL/Decision_Procs/Cooper.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Decision_Procs/Cooper.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1400,9 +1400,8 @@
   also have "\<dots> = (j dvd (- (c*x - ?e)))"
     by (simp only: dvd_minus_iff)
   also have "\<dots> = (j dvd (c* (- x)) + ?e)"
-    apply (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] diff_minus add_ac minus_add_distrib)
-    apply (simp add: algebra_simps)
-    done
+    by (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] add_ac minus_add_distrib)
+      (simp add: algebra_simps)
   also have "\<dots> = Ifm bbs ((- x)#bs) (Dvd j (CN 0 c e))"
     using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"] by simp
   finally show ?case .
@@ -1413,9 +1412,8 @@
   also have "\<dots> = (j dvd (- (c*x - ?e)))"
     by (simp only: dvd_minus_iff)
   also have "\<dots> = (j dvd (c* (- x)) + ?e)"
-    apply (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] diff_minus add_ac minus_add_distrib)
-    apply (simp add: algebra_simps)
-    done
+    by (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] add_ac minus_add_distrib)
+      (simp add: algebra_simps)
   also have "\<dots> = Ifm bbs ((- x)#bs) (Dvd j (CN 0 c e))"
     using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"] by simp
   finally show ?case by simp
--- a/src/HOL/Decision_Procs/MIR.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Decision_Procs/MIR.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1727,7 +1727,7 @@
   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))" 
       by (simp add: nb Let_def split_def isint_Floor isint_neg)
     have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def)
-    also have "\<dots> = (?I (?l (Lt a)))" apply (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) by (simp add: Ia cp cnz Let_def split_def diff_minus)
+    also have "\<dots> = (?I (?l (Lt a)))" apply (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) by (simp add: Ia cp cnz Let_def split_def)
     finally have ?case using l by simp}
   moreover
   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))" 
@@ -1752,13 +1752,13 @@
   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))" 
       by (simp add: nb Let_def split_def isint_Floor isint_neg)
     have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def)
-    also have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_minus)
+    also have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def)
     finally have ?case using l by simp}
   moreover
   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))" 
       by (simp add: nb Let_def split_def isint_Floor isint_neg)
     have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def)
-    also from cn cnz have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_minus[symmetric] add_ac ,arith)
+    also from cn cnz have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def add_ac, arith)
     finally have ?case using l by simp}
   ultimately show ?case by blast
 next
@@ -1777,13 +1777,13 @@
   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))" 
       by (simp add: nb Let_def split_def isint_Floor isint_neg)
     have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def)
-    also have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_minus)
+    also have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def)
     finally have ?case using l by simp}
   moreover
   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))" 
       by (simp add: nb Let_def split_def isint_Floor isint_neg)
     have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def)
-    also from cn cnz have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_minus[symmetric] add_ac, arith)
+    also from cn cnz have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def add_ac, arith)
     finally have ?case using l by simp}
   ultimately show ?case by blast
 next
@@ -1802,13 +1802,13 @@
   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))" 
       by (simp add: nb Let_def split_def isint_Floor isint_neg)
     have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def)
-    also have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_minus)
+    also have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def)
     finally have ?case using l by simp}
   moreover
   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))" 
       by (simp add: nb Let_def split_def isint_Floor isint_neg)
     have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def)
-    also from cn cnz have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_minus[symmetric] add_ac, arith)
+    also from cn cnz have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def add_ac, arith)
     finally have ?case using l by simp}
   ultimately show ?case by blast
 next
@@ -3125,7 +3125,8 @@
     hence pid: "c*i + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
     with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + ?fe = j1" by auto
     hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - ?N i e + real j1" 
-      by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] algebra_simps)
+      by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff])
+        (simp add: algebra_simps)
     with nob  have ?case by blast }
   ultimately show ?case by blast
 next
@@ -3148,11 +3149,12 @@
     hence pid: "c*i + 1 + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
     with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + 1+ ?fe = j1" by auto
     hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) + 1= - ?N i e + real j1"
-      by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] algebra_simps real_of_one) 
+      by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] real_of_one) 
+        (simp add: algebra_simps)
     hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = (- ?N i e + real j1) - 1"
       by (simp only: algebra_simps)
         hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - 1 - ?N i e + real j1"
-          by (simp only: add_ac diff_minus)
+          by (simp add: algebra_simps minus_one [symmetric] del: minus_one)
     with nob  have ?case by blast }
   ultimately show ?case by blast
 next
@@ -3477,10 +3479,7 @@
   qed
 next
   case (3 a b) then show ?case
-    apply auto
-    apply (erule_tac x = "(aa, aaa, ba)" in ballE) apply simp_all
-    apply (erule_tac x = "(ab, ac, baa)" in ballE) apply simp_all
-    done
+    by auto
 qed (auto simp add: Let_def split_def algebra_simps)
 
 lemma real_in_int_intervals: 
@@ -3615,7 +3614,7 @@
       by(simp only: myle[of _ "real n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real n *x + ?N s - ?N (Floor s)"]) 
     hence "\<exists> j\<in> {n .. 0}. 0 \<ge> - (real n *x + ?N s - ?N (Floor s) - real j) \<and> - (real n *x + ?N s - ?N (Floor s) - real (j+1)) > 0" by (simp only: th1[rule_format] th2[rule_format])
     hence "\<exists> j\<in> {n.. 0}. ?I (?p (p,n,s) j)"
-      using pns by (simp add: fp_def nn diff_minus add_ac mult_ac
+      using pns by (simp add: fp_def nn algebra_simps
         del: diff_less_0_iff_less diff_le_0_iff_le) 
     then obtain "j" where j_def: "j\<in> {n .. 0} \<and> ?I (?p (p,n,s) j)" by blast
     hence "\<exists>x \<in> {?p (p,n,s) j |j. n\<le> j \<and> j \<le> 0 }. ?I x" by auto
@@ -4832,7 +4831,7 @@
   shows "(Ifm bs (E p)) = (\<exists> (i::int). Ifm (real i#bs) (E (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (exsplit p))))" (is "?lhs = ?rhs")
 proof-
   have "?rhs = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm (x#(real i)#bs) (exsplit p))"
-    by (simp add: myless[of _ "1"] myless[of _ "0"] add_ac diff_minus)
+    by (simp add: myless[of _ "1"] myless[of _ "0"] add_ac)
   also have "\<dots> = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm ((real i + x) #bs) p)"
     by (simp only: exsplit[OF qf] add_ac)
   also have "\<dots> = (\<exists> x. Ifm (x#bs) p)" 
@@ -5196,7 +5195,7 @@
   hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)" 
     by (auto simp only: subst0_bound0[OF qfmq])
   hence th: "\<forall> j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))"
-    by (auto simp add: simpfm_bound0)
+    by auto
   from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp 
   from Bn jsnb have "\<forall> (b,j) \<in> set ?bjs. numbound0 (Add b (C j))"
     by simp
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1959,7 +1959,7 @@
       by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Eq (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r = 0" 
-      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
+      by (simp add: r [of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
     also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) =0 "
       using c d mult_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r =0" 
@@ -2041,7 +2041,7 @@
       by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (NEq (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r \<noteq> 0" 
-      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
+      by (simp add: r [of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
     also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) \<noteq> 0 "
       using c d mult_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r \<noteq> 0" 
@@ -2106,7 +2106,7 @@
       by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r < 0" 
-      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
+      by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
     also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) < 0"
       
       using dc' dc'' mult_less_cancel_left_disj[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
@@ -2127,7 +2127,7 @@
       by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r < 0" 
-      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
+      by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
 
     also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) > 0"
       
@@ -2251,7 +2251,7 @@
       by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r <= 0" 
-      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
+      by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
     also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) <= 0"
       
       using dc' dc'' mult_le_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
@@ -2272,7 +2272,7 @@
       by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r <= 0" 
-      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
+      by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
 
     also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) >= 0"
       
@@ -2356,8 +2356,11 @@
 
 lemma msubst_I: assumes lp: "islin p" and nc: "isnpoly c" and nd: "isnpoly d"
   shows "Ifm vs (x#bs) (msubst p ((c,t),(d,s))) = Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /2)#bs) p"
-  using lp
-by (induct p rule: islin.induct, auto simp add: tmbound0_I[where b="(- (Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>) + - (Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>)) / 2" and b'=x and bs = bs and vs=vs] bound0_I[where b="(- (Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>) + - (Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>)) / 2" and b'=x and bs = bs and vs=vs] msubsteq msubstneq msubstlt[OF nc nd] msubstle[OF nc nd])
+  using lp by (induct p rule: islin.induct)
+    (auto simp add: tmbound0_I
+    [where b = "(- (Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>) - (Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>)) / 2"
+      and b' = x and bs = bs and vs = vs]
+    msubsteq msubstneq msubstlt [OF nc nd] msubstle [OF nc nd])
 
 lemma msubst_nb: assumes lp: "islin p" and t: "tmbound0 t" and s: "tmbound0 s"
   shows "bound0 (msubst p ((c,t),(d,s)))"
@@ -2429,7 +2432,7 @@
   with evaldjf_bound0[of ?Up "(simpfm o (msubst (simpfm p)))"]
   have "bound0 (evaldjf (simpfm o (msubst (simpfm p))) ?Up)" by blast
   with mp_nb pp_nb 
-  have th1: "bound0 (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up )))" by (simp add: disj_nb)
+  have th1: "bound0 (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up )))" by simp
   from decr0_qf[OF th1] have thqf: "qfree (ferrack p)" by (simp add: ferrack_def Let_def)
   have "?lhs \<longleftrightarrow> (\<exists>x. Ifm vs (x#bs) ?q)" by simp
   also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> (\<exists>(c, t)\<in>set ?U. \<exists>(d, s)\<in>set ?U. ?I (msubst (simpfm p) ((c, t), d, s)))" using fr_eq_msubst[OF lq, of vs bs x] by simp
@@ -2612,7 +2615,7 @@
 lemma msubst2_nb: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" and lp: "islin p" and tnb: "tmbound0 t"
   shows "bound0 (msubst2 p c t)"
 using lp tnb
-by (simp add: msubst2_def msubstneg_nb msubstpos_nb conj_nb disj_nb lt_nb simpfm_bound0)
+by (simp add: msubst2_def msubstneg_nb msubstpos_nb lt_nb simpfm_bound0)
 
 lemma mult_minus2_left: "-2 * (x::'a::comm_ring_1) = - (2 * x)"
   by simp
@@ -2666,8 +2669,8 @@
         using H(3) by (auto simp add: msubst2_def lt[OF stupid(1)]  lt[OF stupid(2)] zero_less_mult_iff mult_less_0_iff)
       from msubst2[OF lp nn nn'(1), of x bs ] H(3) nn'
       have "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p"
-        apply (simp add: add_divide_distrib mult_minus2_left)
-        by (simp add: mult_commute)}
+        by (simp add: add_divide_distrib diff_divide_distrib mult_minus2_left mult_commute)
+    }
     moreover
     {fix c t d s assume H: "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)" 
       "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p"
@@ -2675,7 +2678,9 @@
       hence nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)" "\<lparr>(C (-2, 1) *\<^sub>p c*\<^sub>p d)\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
         using H(3,4) by (simp_all add: polymul_norm n2)
       from msubst2[OF lp nn, of x bs ] H(3,4,5) 
-      have "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))" apply (simp add: add_divide_distrib mult_minus2_left) by (simp add: mult_commute)}
+      have "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))"
+        by (simp add: diff_divide_distrib add_divide_distrib mult_minus2_left mult_commute)
+    }
     ultimately show ?thesis by blast
   qed
   from fr_eq2[OF lp, of vs bs x] show ?thesis
--- a/src/HOL/Decision_Procs/Polynomial_List.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Decision_Procs/Polynomial_List.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -2,371 +2,379 @@
     Author:     Amine Chaieb
 *)
 
-header {* Univariate Polynomials as Lists *}
+header {* Univariate Polynomials as lists *}
 
 theory Polynomial_List
-imports Main
+imports Complex_Main
 begin
 
-text{* Application of polynomial as a real function. *}
+text{* Application of polynomial as a function. *}
 
-primrec poly :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a::comm_ring"
+primrec (in semiring_0) poly :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a"
 where
   poly_Nil:  "poly [] x = 0"
-| poly_Cons: "poly (h # t) x = h + x * poly t x"
+| poly_Cons: "poly (h#t) x = h + x * poly t x"
 
 
 subsection{*Arithmetic Operations on Polynomials*}
 
 text{*addition*}
-primrec padd :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a::comm_ring_1 list"  (infixl "+++" 65)
+
+primrec (in semiring_0) padd :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "+++" 65)
 where
   padd_Nil:  "[] +++ l2 = l2"
-| padd_Cons: "(h # t) +++ l2 = (if l2 = [] then h # t else (h + hd l2) # (t +++ tl l2))"
+| padd_Cons: "(h#t) +++ l2 = (if l2 = [] then h#t else (h + hd l2)#(t +++ tl l2))"
 
 text{*Multiplication by a constant*}
-primrec cmult :: "'a::comm_ring_1 \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "%*" 70)
-where
+primrec (in semiring_0) cmult :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "%*" 70) where
   cmult_Nil:  "c %* [] = []"
-| cmult_Cons: "c %* (h # t) = (c * h) # (c %* t)"
+| cmult_Cons: "c %* (h#t) = (c * h)#(c %* t)"
 
 text{*Multiplication by a polynomial*}
-primrec pmult :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a::comm_ring_1 list"  (infixl "***" 70)
+primrec (in semiring_0) pmult :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "***" 70)
 where
   pmult_Nil:  "[] *** l2 = []"
-| pmult_Cons: "(h # t) *** l2 =
-    (if t = [] then h %* l2 else (h %* l2) +++ (0 # (t *** l2)))"
+| pmult_Cons: "(h#t) *** l2 = (if t = [] then h %* l2
+                              else (h %* l2) +++ ((0) # (t *** l2)))"
 
 text{*Repeated multiplication by a polynomial*}
-primrec mulexp :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a::comm_ring_1 list"
-where
+primrec (in semiring_0) mulexp :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a  list \<Rightarrow> 'a list" where
   mulexp_zero:  "mulexp 0 p q = q"
 | mulexp_Suc:   "mulexp (Suc n) p q = p *** mulexp n p q"
 
 text{*Exponential*}
-primrec pexp :: "'a list \<Rightarrow> nat \<Rightarrow> 'a::comm_ring_1 list"  (infixl "%^" 80)
-where
+primrec (in semiring_1) pexp :: "'a list \<Rightarrow> nat \<Rightarrow> 'a list"  (infixl "%^" 80) where
   pexp_0:   "p %^ 0 = [1]"
 | pexp_Suc: "p %^ (Suc n) = p *** (p %^ n)"
 
 text{*Quotient related value of dividing a polynomial by x + a*}
 (* Useful for divisor properties in inductive proofs *)
-primrec pquot :: "'a list \<Rightarrow> 'a::field \<Rightarrow> 'a list"
+primrec (in field) "pquot" :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list"
 where
-  pquot_Nil: "pquot [] a = []"
-| pquot_Cons: "pquot (h # t) a =
-    (if t = [] then [h] else (inverse a * (h - hd (pquot t a))) # pquot t a)"
-
+  pquot_Nil:  "pquot [] a= []"
+| pquot_Cons: "pquot (h#t) a =
+    (if t = [] then [h] else (inverse(a) * (h - hd( pquot t a)))#(pquot t a))"
 
 text{*normalization of polynomials (remove extra 0 coeff)*}
-primrec pnormalize :: "'a::comm_ring_1 list \<Rightarrow> 'a list"
-where
+primrec (in semiring_0) pnormalize :: "'a list \<Rightarrow> 'a list" where
   pnormalize_Nil:  "pnormalize [] = []"
-| pnormalize_Cons: "pnormalize (h # p) =
-    (if (pnormalize p = []) then (if h = 0 then [] else [h])
-     else (h # pnormalize p))"
+| pnormalize_Cons: "pnormalize (h#p) =
+    (if pnormalize p = [] then (if h = 0 then [] else [h]) else h # pnormalize p)"
 
-definition "pnormal p = ((pnormalize p = p) \<and> p \<noteq> [])"
-definition "nonconstant p = (pnormal p \<and> (\<forall>x. p \<noteq> [x]))"
+definition (in semiring_0) "pnormal p = ((pnormalize p = p) \<and> p \<noteq> [])"
+definition (in semiring_0) "nonconstant p = (pnormal p \<and> (\<forall>x. p \<noteq> [x]))"
 text{*Other definitions*}
 
-definition poly_minus :: "'a list => ('a :: comm_ring_1) list"  ("-- _" [80] 80)
+definition (in ring_1) poly_minus :: "'a list \<Rightarrow> 'a list" ("-- _" [80] 80)
   where "-- p = (- 1) %* p"
 
-definition divides :: "'a::comm_ring_1 list \<Rightarrow> 'a list \<Rightarrow> bool"  (infixl "divides" 70)
-  where "p1 divides p2 = (\<exists>q. poly p2 = poly (p1 *** q))"
+definition (in semiring_0) divides :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"  (infixl "divides" 70)
+  where "p1 divides p2 = (\<exists>q. poly p2 = poly(p1 *** q))"
+
+lemma (in semiring_0) dividesI:
+  "poly p2 = poly (p1 *** q) \<Longrightarrow> p1 divides p2"
+  by (auto simp add: divides_def)
 
-definition order :: "'a::comm_ring_1 \<Rightarrow> 'a list \<Rightarrow> nat" --{*order of a polynomial*}
-  where "order a p = (SOME n. ([-a, 1] %^ n) divides p & ~ (([-a, 1] %^ Suc n) divides p))"
+lemma (in semiring_0) dividesE:
+  assumes "p1 divides p2"
+  obtains q where "poly p2 = poly (p1 *** q)"
+  using assms by (auto simp add: divides_def)
 
-definition degree :: "'a::comm_ring_1 list \<Rightarrow> nat" --{*degree of a polynomial*}
+    --{*order of a polynomial*}
+definition (in ring_1) order :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where
+  "order a p = (SOME n. ([-a, 1] %^ n) divides p \<and> ~ (([-a, 1] %^ (Suc n)) divides p))"
+
+     --{*degree of a polynomial*}
+definition (in semiring_0) degree :: "'a list \<Rightarrow> nat"
   where "degree p = length (pnormalize p) - 1"
 
-definition rsquarefree :: "'a::comm_ring_1 list \<Rightarrow> bool"
-  where --{*squarefree polynomials --- NB with respect to real roots only.*}
-  "rsquarefree p = (poly p \<noteq> poly [] \<and> (\<forall>a. order a p = 0 \<or> order a p = 1))"
+     --{*squarefree polynomials --- NB with respect to real roots only.*}
+definition (in ring_1) rsquarefree :: "'a list \<Rightarrow> bool"
+  where "rsquarefree p \<longleftrightarrow> poly p \<noteq> poly [] \<and> (\<forall>a. order a p = 0 \<or> order a p = 1)"
 
-lemma padd_Nil2 [simp]: "p +++ [] = p"
+context semiring_0
+begin
+
+lemma padd_Nil2[simp]: "p +++ [] = p"
   by (induct p) auto
 
 lemma padd_Cons_Cons: "(h1 # p1) +++ (h2 # p2) = (h1 + h2) # (p1 +++ p2)"
   by auto
 
-lemma pminus_Nil [simp]: "-- [] = []"
+lemma pminus_Nil: "-- [] = []"
   by (simp add: poly_minus_def)
 
-lemma pmult_singleton: "[h1] *** p1 = h1 %* p1"
-  by simp
+lemma pmult_singleton: "[h1] *** p1 = h1 %* p1" by simp
+
+end
 
-lemma poly_ident_mult [simp]: "1 %* t = t"
-  by (induct t) auto
+lemma (in semiring_1) poly_ident_mult[simp]: "1 %* t = t" by (induct t) auto
 
-lemma poly_simple_add_Cons [simp]: "[a] +++ ((0)#t) = (a#t)"
+lemma (in semiring_0) poly_simple_add_Cons[simp]: "[a] +++ ((0)#t) = (a#t)"
   by simp
 
 text{*Handy general properties*}
 
-lemma padd_commut: "b +++ a = a +++ b"
-  apply (induct b arbitrary: a)
-  apply auto
-  apply (rule padd_Cons [THEN ssubst])
-  apply (case_tac aa)
-  apply auto
+lemma (in comm_semiring_0) padd_commut: "b +++ a = a +++ b"
+proof (induct b arbitrary: a)
+  case Nil
+  thus ?case by auto
+next
+  case (Cons b bs a)
+  thus ?case by (cases a) (simp_all add: add_commute)
+qed
+
+lemma (in comm_semiring_0) padd_assoc: "\<forall>b c. (a +++ b) +++ c = a +++ (b +++ c)"
+  apply (induct a)
+  apply (simp, clarify)
+  apply (case_tac b, simp_all add: add_ac)
   done
 
-lemma padd_assoc: "(a +++ b) +++ c = a +++ (b +++ c)"
-  apply (induct a arbitrary: b c)
+lemma (in semiring_0) poly_cmult_distr: "a %* ( p +++ q) = (a %* p +++ a %* q)"
+  apply (induct p arbitrary: q)
   apply simp
-  apply (case_tac b)
-  apply simp_all
+  apply (case_tac q, simp_all add: distrib_left)
   done
 
-lemma poly_cmult_distr: "a %* ( p +++ q) = (a %* p +++ a %* q)"
-  apply (induct p arbitrary: q)
+lemma (in ring_1) pmult_by_x[simp]: "[0, 1] *** t = ((0)#t)"
+  apply (induct t)
   apply simp
-  apply (case_tac q)
-  apply (simp_all add: distrib_left)
+  apply (auto simp add: padd_commut)
+  apply (case_tac t, auto)
   done
 
-lemma pmult_by_x [simp]: "[0, 1] *** t = ((0)#t)"
-  by (induct t) (auto simp add: padd_commut)
-
-
 text{*properties of evaluation of polynomials.*}
 
-lemma poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x"
-  apply (induct p1 arbitrary: p2)
-  apply auto
-  apply (case_tac "p2")
-  apply (auto simp add: distrib_left)
-  done
+lemma (in semiring_0) poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x"
+proof(induct p1 arbitrary: p2)
+  case Nil
+  thus ?case by simp
+next
+  case (Cons a as p2)
+  thus ?case
+    by (cases p2) (simp_all  add: add_ac distrib_left)
+qed
 
-lemma poly_cmult: "poly (c %* p) x = c * poly p x"
+lemma (in comm_semiring_0) poly_cmult: "poly (c %* p) x = c * poly p x"
   apply (induct p)
-  apply simp
-  apply (cases "x = 0")
+  apply (case_tac [2] "x = zero")
   apply (auto simp add: distrib_left mult_ac)
   done
 
-lemma poly_minus: "poly (-- p) x = - (poly p x)"
-  by (simp add: poly_minus_def poly_cmult)
+lemma (in comm_semiring_0) poly_cmult_map: "poly (map (op * c) p) x = c*poly p x"
+  by (induct p) (auto simp add: distrib_left mult_ac)
 
-lemma poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x"
-  apply (induct p1 arbitrary: p2)
-  apply (case_tac p1)
-  apply (auto simp add: poly_cmult poly_add distrib_right distrib_left mult_ac)
+lemma (in comm_ring_1) poly_minus: "poly (-- p) x = - (poly p x)"
+  apply (simp add: poly_minus_def)
+  apply (auto simp add: poly_cmult)
   done
 
-lemma poly_exp: "poly (p %^ n) (x::'a::comm_ring_1) = (poly p x) ^ n"
+lemma (in comm_semiring_0) poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x"
+proof (induct p1 arbitrary: p2)
+  case Nil
+  thus ?case by simp
+next
+  case (Cons a as p2)
+  thus ?case by (cases as)
+    (simp_all add: poly_cmult poly_add distrib_right distrib_left mult_ac)
+qed
+
+class idom_char_0 = idom + ring_char_0
+
+subclass (in field_char_0) idom_char_0 ..
+
+lemma (in comm_ring_1) poly_exp: "poly (p %^ n) x = (poly p x) ^ n"
   by (induct n) (auto simp add: poly_cmult poly_mult)
 
 text{*More Polynomial Evaluation Lemmas*}
 
-lemma poly_add_rzero [simp]: "poly (a +++ []) x = poly a x"
+lemma (in semiring_0) poly_add_rzero[simp]: "poly (a +++ []) x = poly a x"
   by simp
 
-lemma poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x"
+lemma (in comm_semiring_0) poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x"
   by (simp add: poly_mult mult_assoc)
 
-lemma poly_mult_Nil2 [simp]: "poly (p *** []) x = 0"
+lemma (in semiring_0) poly_mult_Nil2[simp]: "poly (p *** []) x = 0"
   by (induct p) auto
 
-lemma poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x"
+lemma (in comm_semiring_1) poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x"
   by (induct n) (auto simp add: poly_mult mult_assoc)
 
 subsection{*Key Property: if @{term "f(a) = 0"} then @{term "(x - a)"} divides
  @{term "p(x)"} *}
 
-lemma poly_linear_rem: "\<exists>q r. h # t = [r] +++ [-a, 1] *** q"
-  apply (induct t arbitrary: h)
-  apply (rule_tac x = "[]" in exI)
-  apply (rule_tac x = h in exI)
-  apply simp
-  apply (drule_tac x = aa in meta_spec)
-  apply safe
-  apply (rule_tac x = "r#q" in exI)
-  apply (rule_tac x = "a*r + h" in exI)
-  apply (case_tac q)
-  apply auto
-  done
+lemma (in comm_ring_1) lemma_poly_linear_rem: "\<forall>h. \<exists>q r. h#t = [r] +++ [-a, 1] *** q"
+proof(induct t)
+  case Nil
+  { fix h have "[h] = [h] +++ [- a, 1] *** []" by simp }
+  thus ?case by blast
+next
+  case (Cons  x xs)
+  { fix h
+    from Cons.hyps[rule_format, of x]
+    obtain q r where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast
+    have "h#x#xs = [a*r + h] +++ [-a, 1] *** (r#q)"
+      using qr by (cases q) (simp_all add: algebra_simps)
+    hence "\<exists>q r. h#x#xs = [r] +++ [-a, 1] *** q" by blast}
+  thus ?case by blast
+qed
+
+lemma (in comm_ring_1) poly_linear_rem: "\<exists>q r. h#t = [r] +++ [-a, 1] *** q"
+  using lemma_poly_linear_rem [where t = t and a = a] by auto
+
 
-lemma poly_linear_divides: "poly p a = 0 \<longleftrightarrow> p = [] \<or> (\<exists>q. p = [-a, 1] *** q)"
-  apply (auto simp add: poly_add poly_cmult distrib_left)
-  apply (case_tac p)
-  apply simp
-  apply (cut_tac h = aa and t = list and a = a in poly_linear_rem)
-  apply safe
-  apply (case_tac q)
-  apply auto
-  apply (drule_tac x = "[]" in spec)
-  apply simp
-  apply (auto simp add: poly_add poly_cmult add_assoc)
-  apply (drule_tac x = "aa#lista" in spec)
-  apply auto
-  done
+lemma (in comm_ring_1) poly_linear_divides: "(poly p a = 0) = ((p = []) | (\<exists>q. p = [-a, 1] *** q))"
+proof -
+  { assume p: "p = []" hence ?thesis by simp }
+  moreover
+  {
+    fix x xs assume p: "p = x#xs"
+    {
+      fix q assume "p = [-a, 1] *** q"
+      hence "poly p a = 0" by (simp add: poly_add poly_cmult)
+    }
+    moreover
+    { assume p0: "poly p a = 0"
+      from poly_linear_rem[of x xs a] obtain q r
+      where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast
+      have "r = 0" using p0 by (simp only: p qr poly_mult poly_add) simp
+      hence "\<exists>q. p = [- a, 1] *** q"
+        using p qr
+        apply -
+        apply (rule exI[where x=q])
+        apply auto
+        apply (cases q)
+        apply auto
+        done
+    }
+    ultimately have ?thesis using p by blast
+  }
+  ultimately show ?thesis by (cases p) auto
+qed
 
-lemma lemma_poly_length_mult [simp]: "length (k %* p +++  (h # (a %* p))) = Suc (length p)"
-  by (induct p arbitrary: h k a) auto
+lemma (in semiring_0) lemma_poly_length_mult[simp]: "\<forall>h k a. length (k %* p +++  (h # (a %* p))) = Suc (length p)"
+  by (induct p) auto
 
-lemma lemma_poly_length_mult2 [simp]: "length (k %* p +++  (h # p)) = Suc (length p)"
-  by (induct p arbitrary: h k) auto
+lemma (in semiring_0) lemma_poly_length_mult2[simp]: "\<forall>h k. length (k %* p +++  (h # p)) = Suc (length p)"
+  by (induct p) auto
 
-lemma poly_length_mult [simp]: "length([-a, 1] *** q) = Suc (length q)"
+lemma (in ring_1) poly_length_mult[simp]: "length([-a,1] *** q) = Suc (length q)"
   by auto
 
-
 subsection{*Polynomial length*}
 
-lemma poly_cmult_length [simp]: "length (a %* p) = length p"
+lemma (in semiring_0) poly_cmult_length[simp]: "length (a %* p) = length p"
   by (induct p) auto
 
-lemma poly_add_length:
-  "length (p1 +++ p2) = (if (length p1 < length p2) then length p2 else length p1)"
-  by (induct p1 arbitrary: p2) auto
+lemma (in semiring_0) poly_add_length: "length (p1 +++ p2) = max (length p1) (length p2)"
+  by (induct p1 arbitrary: p2) (simp_all, arith)
 
-lemma poly_root_mult_length [simp]: "length ([a, b] *** p) = Suc (length p)"
-  by simp
+lemma (in semiring_0) poly_root_mult_length[simp]: "length([a,b] *** p) = Suc (length p)"
+  by (simp add: poly_add_length)
 
-lemma poly_mult_not_eq_poly_Nil [simp]:
-  "poly (p *** q) x \<noteq> poly [] x \<longleftrightarrow> poly p x \<noteq> poly [] x \<and> poly q x \<noteq> poly [] (x::'a::idom)"
+lemma (in idom) poly_mult_not_eq_poly_Nil[simp]:
+  "poly (p *** q) x \<noteq> poly [] x \<longleftrightarrow> poly p x \<noteq> poly [] x \<and> poly q x \<noteq> poly [] x"
   by (auto simp add: poly_mult)
 
-lemma poly_mult_eq_zero_disj: "poly (p *** q) (x::'a::idom) = 0 \<longleftrightarrow> poly p x = 0 \<or> poly q x = 0"
+lemma (in idom) poly_mult_eq_zero_disj: "poly (p *** q) x = 0 \<longleftrightarrow> poly p x = 0 \<or> poly q x = 0"
   by (auto simp add: poly_mult)
 
 text{*Normalisation Properties*}
 
-lemma poly_normalized_nil: "pnormalize p = [] \<longrightarrow> poly p x = 0"
+lemma (in semiring_0) poly_normalized_nil: "(pnormalize p = []) --> (poly p x = 0)"
   by (induct p) auto
 
 text{*A nontrivial polynomial of degree n has no more than n roots*}
+lemma (in idom) poly_roots_index_lemma:
+   assumes p: "poly p x \<noteq> poly [] x" and n: "length p = n"
+  shows "\<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)"
+  using p n
+proof (induct n arbitrary: p x)
+  case 0
+  thus ?case by simp
+next
+  case (Suc n p x)
+  {
+    assume C: "\<And>i. \<exists>x. poly p x = 0 \<and> (\<forall>m\<le>Suc n. x \<noteq> i m)"
+    from Suc.prems have p0: "poly p x \<noteq> 0" "p\<noteq> []" by auto
+    from p0(1)[unfolded poly_linear_divides[of p x]]
+    have "\<forall>q. p \<noteq> [- x, 1] *** q" by blast
+    from C obtain a where a: "poly p a = 0" by blast
+    from a[unfolded poly_linear_divides[of p a]] p0(2)
+    obtain q where q: "p = [-a, 1] *** q" by blast
+    have lg: "length q = n" using q Suc.prems(2) by simp
+    from q p0 have qx: "poly q x \<noteq> poly [] x"
+      by (auto simp add: poly_mult poly_add poly_cmult)
+    from Suc.hyps[OF qx lg] obtain i where
+      i: "\<forall>x. poly q x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)" by blast
+    let ?i = "\<lambda>m. if m = Suc n then a else i m"
+    from C[of ?i] obtain y where y: "poly p y = 0" "\<forall>m\<le> Suc n. y \<noteq> ?i m"
+      by blast
+    from y have "y = a \<or> poly q y = 0"
+      by (simp only: q poly_mult_eq_zero_disj poly_add) (simp add: algebra_simps)
+    with i[rule_format, of y] y(1) y(2) have False
+      apply auto
+      apply (erule_tac x = "m" in allE)
+      apply auto
+      done
+  }
+  thus ?case by blast
+qed
 
-lemma poly_roots_index_lemma0 [rule_format]:
-   "\<forall>p x. poly p x \<noteq> poly [] x & length p = n
-    --> (\<exists>i. \<forall>x. (poly p x = (0::'a::idom)) --> (\<exists>m. (m \<le> n & x = i m)))"
-  apply (induct n)
-  apply safe
-  apply (rule ccontr)
-  apply (subgoal_tac "\<exists>a. poly p a = 0")
-  apply safe
-  apply (drule poly_linear_divides [THEN iffD1])
-  apply safe
-  apply (drule_tac x = q in spec)
-  apply (drule_tac x = x in spec)
-  apply (simp del: poly_Nil pmult_Cons)
-  apply (erule exE)
-  apply (drule_tac x = "%m. if m = Suc n then a else i m" in spec)
-  apply safe
-  apply (drule poly_mult_eq_zero_disj [THEN iffD1])
-  apply safe
-  apply (drule_tac x = "Suc (length q)" in spec)
-  apply (auto simp add: field_simps)
-  apply (drule_tac x = xa in spec)
-  apply (clarsimp simp add: field_simps)
-  apply (drule_tac x = m in spec)
-  apply (auto simp add:field_simps)
-  done
-lemmas poly_roots_index_lemma1 = conjI [THEN poly_roots_index_lemma0]
 
-lemma poly_roots_index_length0:
-  "poly p (x::'a::idom) \<noteq> poly [] x \<Longrightarrow>
-    \<exists>i. \<forall>x. (poly p x = 0) \<longrightarrow> (\<exists>n. n \<le> length p & x = i n)"
-  by (blast intro: poly_roots_index_lemma1)
+lemma (in idom) poly_roots_index_length:
+  "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>i. \<forall>x. (poly p x = 0) \<longrightarrow> (\<exists>n. n \<le> length p \<and> x = i n)"
+  by (blast intro: poly_roots_index_lemma)
 
-lemma poly_roots_finite_lemma:
-  "poly p (x::'a::idom) \<noteq> poly [] x \<Longrightarrow>
-    \<exists>N i. \<forall>x. (poly p x = 0) \<longrightarrow> (\<exists>n. (n::nat) < N & x = i n)"
-  apply (drule poly_roots_index_length0)
-  apply safe
+lemma (in idom) poly_roots_finite_lemma1:
+  "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>N i. \<forall>x. (poly p x = 0) \<longrightarrow> (\<exists>n. (n::nat) < N \<and> x = i n)"
+  apply (drule poly_roots_index_length, safe)
   apply (rule_tac x = "Suc (length p)" in exI)
   apply (rule_tac x = i in exI)
   apply (simp add: less_Suc_eq_le)
   done
 
-
-lemma real_finite_lemma:
-  assumes "\<forall>x. P x \<longrightarrow> (\<exists>n. n < length j & x = j!n)"
-  shows "finite {(x::'a::idom). P x}"
+lemma (in idom) idom_finite_lemma:
+  assumes P: "\<forall>x. P x --> (\<exists>n. n < length j \<and> x = j!n)"
+  shows "finite {x. P x}"
 proof -
   let ?M = "{x. P x}"
   let ?N = "set j"
-  have "?M \<subseteq> ?N" using assms by auto
-  then show ?thesis using finite_subset by auto
+  have "?M \<subseteq> ?N" using P by auto
+  thus ?thesis using finite_subset by auto
 qed
 
-lemma poly_roots_index_lemma [rule_format]:
-  "\<forall>p x. poly p x \<noteq> poly [] x & length p = n
-    \<longrightarrow> (\<exists>i. \<forall>x. (poly p x = (0::'a::{idom})) \<longrightarrow> x \<in> set i)"
-  apply (induct n)
-  apply safe
-  apply (rule ccontr)
-  apply (subgoal_tac "\<exists>a. poly p a = 0")
-  apply safe
-  apply (drule poly_linear_divides [THEN iffD1])
-  apply safe
-  apply (drule_tac x = q in spec)
-  apply (drule_tac x = x in spec)
-  apply (auto simp del: poly_Nil pmult_Cons)
-  apply (drule_tac x = "a#i" in spec)
-  apply (auto simp only: poly_mult List.list.size)
-  apply (drule_tac x = xa in spec)
-  apply (clarsimp simp add: field_simps)
-  done
-
-lemmas poly_roots_index_lemma2 = conjI [THEN poly_roots_index_lemma]
-
-lemma poly_roots_index_length:
-  "poly p (x::'a::idom) \<noteq> poly [] x \<Longrightarrow>
-    \<exists>i. \<forall>x. (poly p x = 0) --> x \<in> set i"
-  by (blast intro: poly_roots_index_lemma2)
-
-lemma poly_roots_finite_lemma':
-  "poly p (x::'a::idom) \<noteq> poly [] x \<Longrightarrow>
-    \<exists>i. \<forall>x. (poly p x = 0) --> x \<in> set i"
-  apply (drule poly_roots_index_length)
-  apply auto
+lemma (in idom) poly_roots_finite_lemma2:
+  "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> x \<in> set i"
+  apply (drule poly_roots_index_length, safe)
+  apply (rule_tac x="map (\<lambda>n. i n) [0 ..< Suc (length p)]" in exI)
+  apply (auto simp add: image_iff)
+  apply (erule_tac x="x" in allE, clarsimp)
+  apply (case_tac "n = length p")
+  apply (auto simp add: order_le_less)
   done
 
-lemma UNIV_nat_infinite: "\<not> finite (UNIV :: nat set)"
-  unfolding finite_conv_nat_seg_image
-proof (auto simp add: set_eq_iff image_iff)
-  fix n::nat and f:: "nat \<Rightarrow> nat"
-  let ?N = "{i. i < n}"
-  let ?fN = "f ` ?N"
-  let ?y = "Max ?fN + 1"
-  from nat_seg_image_imp_finite[of "?fN" "f" n]
-  have thfN: "finite ?fN" by simp
-  { assume "n =0" hence "\<exists>x. \<forall>xa<n. x \<noteq> f xa" by auto }
-  moreover
-  { assume nz: "n \<noteq> 0"
-    hence thne: "?fN \<noteq> {}" by auto
-    have "\<forall>x\<in> ?fN. Max ?fN \<ge> x" using nz Max_ge_iff[OF thfN thne] by auto
-    hence "\<forall>x\<in> ?fN. ?y > x" by (auto simp add: less_Suc_eq_le)
-    hence "?y \<notin> ?fN" by auto
-    hence "\<exists>x. \<forall>xa<n. x \<noteq> f xa" by auto }
-  ultimately show "\<exists>x. \<forall>xa<n. x \<noteq> f xa" by blast
+lemma (in ring_char_0) UNIV_ring_char_0_infinte: "\<not> (finite (UNIV:: 'a set))"
+proof
+  assume F: "finite (UNIV :: 'a set)"
+  have "finite (UNIV :: nat set)"
+  proof (rule finite_imageD)
+    have "of_nat ` UNIV \<subseteq> UNIV" by simp
+    then show "finite (of_nat ` UNIV :: 'a set)" using F by (rule finite_subset)
+    show "inj (of_nat :: nat \<Rightarrow> 'a)" by (simp add: inj_on_def)
+  qed
+  with infinite_UNIV_nat show False ..
 qed
 
-lemma UNIV_ring_char_0_infinte: "\<not> finite (UNIV:: ('a::ring_char_0) set)"
+lemma (in idom_char_0) poly_roots_finite: "poly p \<noteq> poly [] \<longleftrightarrow> finite {x. poly p x = 0}"
 proof
-  assume F: "finite (UNIV :: 'a set)"
-  have th0: "of_nat ` UNIV \<subseteq> (UNIV :: 'a set)" by simp
-  from finite_subset[OF th0 F] have th: "finite (of_nat ` UNIV :: 'a set)" .
-  have th': "inj_on (of_nat::nat \<Rightarrow> 'a) UNIV"
-    unfolding inj_on_def by auto
-  from finite_imageD[OF th th'] UNIV_nat_infinite
-  show False by blast
-qed
-
-lemma poly_roots_finite: "poly p \<noteq> poly [] \<longleftrightarrow> finite {x. poly p x = (0::'a::{idom,ring_char_0})}"
-proof
-  assume "poly p \<noteq> poly []"
-  then show "finite {x. poly p x = (0::'a)}"
+  assume H: "poly p \<noteq> poly []"
+  show "finite {x. poly p x = (0::'a)}"
+    using H
     apply -
-    apply (erule contrapos_np)
-    apply (rule ext)
+    apply (erule contrapos_np, rule ext)
     apply (rule ccontr)
-    apply (clarify dest!: poly_roots_finite_lemma')
+    apply (clarify dest!: poly_roots_finite_lemma2)
     using finite_subset
   proof -
     fix x i
@@ -377,119 +385,142 @@
     with finite_subset F show False by auto
   qed
 next
-  assume "finite {x. poly p x = (0\<Colon>'a)}"
-  then show "poly p \<noteq> poly []"
-    using UNIV_ring_char_0_infinte by auto
+  assume F: "finite {x. poly p x = (0\<Colon>'a)}"
+  show "poly p \<noteq> poly []" using F UNIV_ring_char_0_infinte by auto
 qed
 
 text{*Entirety and Cancellation for polynomials*}
 
-lemma poly_entire_lemma:
-  "poly (p:: ('a::{idom,ring_char_0}) list) \<noteq> poly [] \<Longrightarrow> poly q \<noteq> poly [] \<Longrightarrow>
-    poly (p *** q) \<noteq> poly []"
-  by (auto simp add: poly_roots_finite poly_mult Collect_disj_eq)
+lemma (in idom_char_0) poly_entire_lemma2:
+  assumes p0: "poly p \<noteq> poly []"
+    and q0: "poly q \<noteq> poly []"
+  shows "poly (p***q) \<noteq> poly []"
+proof -
+  let ?S = "\<lambda>p. {x. poly p x = 0}"
+  have "?S (p *** q) = ?S p \<union> ?S q" by (auto simp add: poly_mult)
+  with p0 q0 show ?thesis  unfolding poly_roots_finite by auto
+qed
 
-lemma poly_entire:
-  "poly (p *** q) = poly ([]::('a::{idom,ring_char_0}) list) \<longleftrightarrow>
-    (poly p = poly [] \<or> poly q = poly [])"
-  apply (auto dest: fun_cong simp add: poly_entire_lemma poly_mult)
-  apply (blast intro: ccontr dest: poly_entire_lemma poly_mult [THEN subst])
-  done
+lemma (in idom_char_0) poly_entire:
+  "poly (p *** q) = poly [] \<longleftrightarrow> poly p = poly [] \<or> poly q = poly []"
+  using poly_entire_lemma2[of p q]
+  by (auto simp add: fun_eq_iff poly_mult)
 
-lemma poly_entire_neg:
-  "poly (p *** q) \<noteq> poly ([]::('a::{idom,ring_char_0}) list) \<longleftrightarrow>
-    poly p \<noteq> poly [] \<and> poly q \<noteq> poly []"
+lemma (in idom_char_0) poly_entire_neg:
+  "poly (p *** q) \<noteq> poly [] \<longleftrightarrow> poly p \<noteq> poly [] \<and> poly q \<noteq> poly []"
   by (simp add: poly_entire)
 
 lemma fun_eq: "f = g \<longleftrightarrow> (\<forall>x. f x = g x)"
   by auto
 
-lemma poly_add_minus_zero_iff: "poly (p +++ -- q) = poly [] \<longleftrightarrow> poly p = poly q"
-  by (auto simp add: field_simps poly_add poly_minus_def fun_eq poly_cmult)
+lemma (in comm_ring_1) poly_add_minus_zero_iff:
+  "poly (p +++ -- q) = poly [] \<longleftrightarrow> poly p = poly q"
+  by (auto simp add: algebra_simps poly_add poly_minus_def fun_eq poly_cmult)
 
-lemma poly_add_minus_mult_eq: "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))"
-  by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult distrib_left)
+lemma (in comm_ring_1) poly_add_minus_mult_eq:
+  "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))"
+  by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult algebra_simps)
+
+subclass (in idom_char_0) comm_ring_1 ..
 
-lemma poly_mult_left_cancel:
-  "(poly (p *** q) = poly (p *** r)) =
-    (poly p = poly ([]::('a::{idom,ring_char_0}) list) | poly q = poly r)"
-  apply (rule_tac p1 = "p *** q" in poly_add_minus_zero_iff [THEN subst])
-  apply (auto simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff)
-  done
+lemma (in idom_char_0) poly_mult_left_cancel:
+  "poly (p *** q) = poly (p *** r) \<longleftrightarrow> poly p = poly [] \<or> poly q = poly r"
+proof -
+  have "poly (p *** q) = poly (p *** r) \<longleftrightarrow> poly (p *** q +++ -- (p *** r)) = poly []"
+    by (simp only: poly_add_minus_zero_iff)
+  also have "\<dots> \<longleftrightarrow> poly p = poly [] \<or> poly q = poly r"
+    by (auto intro: simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff)
+  finally show ?thesis .
+qed
 
-lemma poly_exp_eq_zero [simp]:
-  "poly (p %^ n) = poly ([]::('a::idom) list) \<longleftrightarrow> poly p = poly [] \<and> n \<noteq> 0"
+lemma (in idom) poly_exp_eq_zero[simp]:
+  "poly (p %^ n) = poly [] \<longleftrightarrow> poly p = poly [] \<and> n \<noteq> 0"
   apply (simp only: fun_eq add: HOL.all_simps [symmetric])
   apply (rule arg_cong [where f = All])
   apply (rule ext)
-  apply (induct_tac n)
-  apply (auto simp add: poly_mult)
+  apply (induct n)
+  apply (auto simp add: poly_exp poly_mult)
   done
 
-lemma poly_prime_eq_zero [simp]: "poly [a, 1::'a::comm_ring_1] \<noteq> poly []"
+lemma (in comm_ring_1) poly_prime_eq_zero[simp]: "poly [a,1] \<noteq> poly []"
   apply (simp add: fun_eq)
-  apply (rule_tac x = "1 - a" in exI)
-  apply simp
+  apply (rule_tac x = "minus one a" in exI)
+  apply (simp add: add_commute [of a])
   done
 
-lemma poly_exp_prime_eq_zero [simp]: "poly ([a, (1::'a::idom)] %^ n) \<noteq> poly []"
+lemma (in idom) poly_exp_prime_eq_zero: "poly ([a, 1] %^ n) \<noteq> poly []"
   by auto
 
 text{*A more constructive notion of polynomials being trivial*}
 
-lemma poly_zero_lemma':
-  "poly (h # t) = poly [] \<Longrightarrow> h = (0::'a::{idom,ring_char_0}) & poly t = poly []"
+lemma (in idom_char_0) poly_zero_lemma': "poly (h # t) = poly [] \<Longrightarrow> h = 0 \<and> poly t = poly []"
   apply (simp add: fun_eq)
-  apply (case_tac "h = 0")
-  apply (drule_tac [2] x = 0 in spec)
-  apply auto
-  apply (case_tac "poly t = poly []")
-  apply simp
+  apply (case_tac "h = zero")
+  apply (drule_tac [2] x = zero in spec, auto)
+  apply (cases "poly t = poly []", simp)
 proof -
   fix x
-  assume H: "\<forall>x. x = (0\<Colon>'a) \<or> poly t x = (0\<Colon>'a)"  and pnz: "poly t \<noteq> poly []"
+  assume H: "\<forall>x. x = (0\<Colon>'a) \<or> poly t x = (0\<Colon>'a)"
+    and pnz: "poly t \<noteq> poly []"
   let ?S = "{x. poly t x = 0}"
   from H have "\<forall>x. x \<noteq>0 \<longrightarrow> poly t x = 0" by blast
   hence th: "?S \<supseteq> UNIV - {0}" by auto
   from poly_roots_finite pnz have th': "finite ?S" by blast
-  from finite_subset[OF th th'] UNIV_ring_char_0_infinte[where ?'a = 'a]
-  show "poly t x = (0\<Colon>'a)" by simp
+  from finite_subset[OF th th'] UNIV_ring_char_0_infinte show "poly t x = (0\<Colon>'a)"
+    by simp
 qed
 
-lemma poly_zero: "poly p = poly [] \<longleftrightarrow> list_all (\<lambda>c. c = (0::'a::{idom,ring_char_0})) p"
+lemma (in idom_char_0) poly_zero: "(poly p = poly []) = list_all (%c. c = 0) p"
   apply (induct p)
   apply simp
   apply (rule iffI)
-  apply (drule poly_zero_lemma')
-  apply auto
+  apply (drule poly_zero_lemma', auto)
   done
 
+lemma (in idom_char_0) poly_0: "list_all (\<lambda>c. c = 0) p \<Longrightarrow> poly p x = 0"
+  unfolding poly_zero[symmetric] by simp
+
+
 
 text{*Basics of divisibility.*}
 
-lemma poly_primes: "[a, (1::'a::idom)] divides (p *** q) \<longleftrightarrow> [a, 1] divides p \<or> [a, 1] divides q"
+lemma (in idom) poly_primes:
+  "[a, 1] divides (p *** q) \<longleftrightarrow> [a, 1] divides p \<or> [a, 1] divides q"
   apply (auto simp add: divides_def fun_eq poly_mult poly_add poly_cmult distrib_right [symmetric])
-  apply (drule_tac x = "-a" in spec)
+  apply (drule_tac x = "uminus a" in spec)
+  apply (simp add: poly_linear_divides poly_add poly_cmult distrib_right [symmetric])
+  apply (cases "p = []")
+  apply (rule exI[where x="[]"])
+  apply simp
+  apply (cases "q = []")
+  apply (erule allE[where x="[]"], simp)
+
+  apply clarsimp
+  apply (cases "\<exists>q\<Colon>'a list. p = a %* q +++ ((0\<Colon>'a) # q)")
+  apply (clarsimp simp add: poly_add poly_cmult)
+  apply (rule_tac x="qa" in exI)
+  apply (simp add: distrib_right [symmetric])
+  apply clarsimp
+
   apply (auto simp add: poly_linear_divides poly_add poly_cmult distrib_right [symmetric])
-  apply (rule_tac x = "qa *** q" in exI)
-  apply (rule_tac [2] x = "p *** qa" in exI)
+  apply (rule_tac x = "pmult qa q" in exI)
+  apply (rule_tac [2] x = "pmult p qa" in exI)
   apply (auto simp add: poly_add poly_mult poly_cmult mult_ac)
   done
 
-lemma poly_divides_refl [simp]: "p divides p"
+lemma (in comm_semiring_1) poly_divides_refl[simp]: "p divides p"
   apply (simp add: divides_def)
-  apply (rule_tac x = "[1]" in exI)
+  apply (rule_tac x = "[one]" in exI)
   apply (auto simp add: poly_mult fun_eq)
   done
 
-lemma poly_divides_trans: "p divides q \<Longrightarrow> q divides r \<Longrightarrow> p divides r"
-  apply (simp add: divides_def)
-  apply safe
-  apply (rule_tac x = "qa *** qaa" in exI)
+lemma (in comm_semiring_1) poly_divides_trans: "p divides q \<Longrightarrow> q divides r \<Longrightarrow> p divides r"
+  apply (simp add: divides_def, safe)
+  apply (rule_tac x = "pmult qa qaa" in exI)
   apply (auto simp add: poly_mult fun_eq mult_assoc)
   done
 
-lemma poly_divides_exp: "m \<le> n \<Longrightarrow> (p %^ m) divides (p %^ n)"
+lemma (in comm_semiring_1) poly_divides_exp: "m \<le> n \<Longrightarrow> (p %^ m) divides (p %^ n)"
   apply (auto simp add: le_iff_add)
   apply (induct_tac k)
   apply (rule_tac [2] poly_divides_trans)
@@ -498,34 +529,37 @@
   apply (auto simp add: poly_mult fun_eq mult_ac)
   done
 
-lemma poly_exp_divides: "(p %^ n) divides q \<Longrightarrow> m \<le> n \<Longrightarrow> (p %^ m) divides q"
+lemma (in comm_semiring_1) poly_exp_divides:
+  "(p %^ n) divides q \<Longrightarrow> m \<le> n \<Longrightarrow> (p %^ m) divides q"
   by (blast intro: poly_divides_exp poly_divides_trans)
 
-lemma poly_divides_add: "p divides q \<Longrightarrow> p divides r \<Longrightarrow> p divides (q +++ r)"
-  apply (simp add: divides_def)
-  apply auto
-  apply (rule_tac x = "qa +++ qaa" in exI)
+lemma (in comm_semiring_0) poly_divides_add:
+  "p divides q \<Longrightarrow> p divides r \<Longrightarrow> p divides (q +++ r)"
+  apply (simp add: divides_def, auto)
+  apply (rule_tac x = "padd qa qaa" in exI)
   apply (auto simp add: poly_add fun_eq poly_mult distrib_left)
   done
 
-lemma poly_divides_diff: "p divides q \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides r"
-  apply (auto simp add: divides_def)
-  apply (rule_tac x = "qaa +++ -- qa" in exI)
+lemma (in comm_ring_1) poly_divides_diff:
+  "p divides q \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides r"
+  apply (simp add: divides_def, auto)
+  apply (rule_tac x = "padd qaa (poly_minus qa)" in exI)
   apply (auto simp add: poly_add fun_eq poly_mult poly_minus algebra_simps)
   done
 
-lemma poly_divides_diff2: "p divides r \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides q"
+lemma (in comm_ring_1) poly_divides_diff2:
+  "p divides r \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides q"
   apply (erule poly_divides_diff)
   apply (auto simp add: poly_add fun_eq poly_mult divides_def add_ac)
   done
 
-lemma poly_divides_zero: "poly p = poly [] \<Longrightarrow> q divides p"
+lemma (in semiring_0) poly_divides_zero: "poly p = poly [] \<Longrightarrow> q divides p"
   apply (simp add: divides_def)
-  apply (rule exI [where x = "[]"])
+  apply (rule exI[where x="[]"])
   apply (auto simp add: fun_eq poly_mult)
   done
 
-lemma poly_divides_zero2 [simp]: "q divides []"
+lemma (in semiring_0) poly_divides_zero2 [simp]: "q divides []"
   apply (simp add: divides_def)
   apply (rule_tac x = "[]" in exI)
   apply (auto simp add: fun_eq)
@@ -533,195 +567,256 @@
 
 text{*At last, we can consider the order of a root.*}
 
-lemma poly_order_exists_lemma [rule_format]:
-  "\<forall>p. length p = d \<longrightarrow> poly p \<noteq> poly [] \<longrightarrow>
-    (\<exists>n q. p = mulexp n [-a, (1::'a::{idom,ring_char_0})] q & poly q a \<noteq> 0)"
-  apply (induct "d")
-  apply (simp add: fun_eq)
-  apply safe
-  apply (case_tac "poly p a = 0")
-  apply (drule_tac poly_linear_divides [THEN iffD1])
-  apply safe
-  apply (drule_tac x = q in spec)
-  apply (drule_tac poly_entire_neg [THEN iffD1])
-  apply safe
-  apply force
-  apply (rule_tac x = "Suc n" in exI)
-  apply (rule_tac x = qa in exI)
-  apply (simp del: pmult_Cons)
-  apply (rule_tac x = 0 in exI)
-  apply force
-  done
+lemma (in idom_char_0) poly_order_exists_lemma:
+  assumes lp: "length p = d"
+    and p: "poly p \<noteq> poly []"
+  shows "\<exists>n q. p = mulexp n [-a, 1] q \<and> poly q a \<noteq> 0"
+  using lp p
+proof (induct d arbitrary: p)
+  case 0
+  thus ?case by simp
+next
+  case (Suc n p)
+  show ?case
+  proof (cases "poly p a = 0")
+    case True
+    from Suc.prems have h: "length p = Suc n" "poly p \<noteq> poly []" by auto
+    hence pN: "p \<noteq> []" by auto
+    from True[unfolded poly_linear_divides] pN obtain q where q: "p = [-a, 1] *** q"
+      by blast
+    from q h True have qh: "length q = n" "poly q \<noteq> poly []"
+      apply -
+      apply simp
+      apply (simp only: fun_eq)
+      apply (rule ccontr)
+      apply (simp add: fun_eq poly_add poly_cmult)
+      done
+    from Suc.hyps[OF qh] obtain m r where mr: "q = mulexp m [-a,1] r" "poly r a \<noteq> 0"
+      by blast
+    from mr q have "p = mulexp (Suc m) [-a,1] r \<and> poly r a \<noteq> 0" by simp
+    then show ?thesis by blast
+  next
+    case False
+    then show ?thesis
+      using Suc.prems
+      apply simp
+      apply (rule exI[where x="0::nat"])
+      apply simp
+      done
+  qed
+qed
+
+
+lemma (in comm_semiring_1) poly_mulexp: "poly (mulexp n p q) x = (poly p x) ^ n * poly q x"
+  by (induct n) (auto simp add: poly_mult mult_ac)
+
+lemma (in comm_semiring_1) divides_left_mult:
+  assumes d:"(p***q) divides r" shows "p divides r \<and> q divides r"
+proof-
+  from d obtain t where r:"poly r = poly (p***q *** t)"
+    unfolding divides_def by blast
+  hence "poly r = poly (p *** (q *** t))"
+    "poly r = poly (q *** (p***t))" by(auto simp add: fun_eq poly_mult mult_ac)
+  thus ?thesis unfolding divides_def by blast
+qed
+
 
 (* FIXME: Tidy up *)
-lemma poly_order_exists:
-  "length p = d \<Longrightarrow> poly p \<noteq> poly [] \<Longrightarrow>
-    \<exists>n. ([-a, 1] %^ n) divides p \<and> \<not> (([-a, (1::'a::{idom,ring_char_0})] %^ (Suc n)) divides p)"
-  apply (drule poly_order_exists_lemma [where a=a])
-  apply assumption
-  apply clarify
-  apply (rule_tac x = n in exI)
-  apply safe
-  apply (unfold divides_def)
-  apply (rule_tac x = q in exI)
-  apply (induct_tac n)
-  apply simp
-  apply (simp add: poly_add poly_cmult poly_mult distrib_left mult_ac)
-  apply safe
-  apply (subgoal_tac "poly (mulexp n [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc n *** qa)")
-  apply simp
-  apply (induct_tac n)
-  apply (simp del: pmult_Cons pexp_Suc)
-  apply (erule_tac Q = "poly q a = 0" in contrapos_np)
-  apply (simp add: poly_add poly_cmult)
-  apply (rule pexp_Suc [THEN ssubst])
-  apply (rule ccontr)
-  apply (simp add: poly_mult_left_cancel poly_mult_assoc del: pmult_Cons pexp_Suc)
-  done
+
+lemma (in semiring_1) zero_power_iff: "0 ^ n = (if n = 0 then 1 else 0)"
+  by (induct n) simp_all
 
-lemma poly_one_divides [simp]: "[1] divides p"
-  by (auto simp: divides_def)
+lemma (in idom_char_0) poly_order_exists:
+  assumes "length p = d" and "poly p \<noteq> poly []"
+  shows "\<exists>n. [- a, 1] %^ n divides p \<and> \<not> [- a, 1] %^ Suc n divides p"
+proof -
+  from assms have "\<exists>n q. p = mulexp n [- a, 1] q \<and> poly q a \<noteq> 0"
+    by (rule poly_order_exists_lemma)
+  then obtain n q where p: "p = mulexp n [- a, 1] q" and "poly q a \<noteq> 0" by blast
+  have "[- a, 1] %^ n divides mulexp n [- a, 1] q"
+  proof (rule dividesI)
+    show "poly (mulexp n [- a, 1] q) = poly ([- a, 1] %^ n *** q)"
+      by (induct n) (simp_all add: poly_add poly_cmult poly_mult algebra_simps)
+  qed
+  moreover have "\<not> [- a, 1] %^ Suc n divides mulexp n [- a, 1] q"
+  proof
+    assume "[- a, 1] %^ Suc n divides mulexp n [- a, 1] q"
+    then obtain m where "poly (mulexp n [- a, 1] q) = poly ([- a, 1] %^ Suc n *** m)"
+      by (rule dividesE)
+    moreover have "poly (mulexp n [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc n *** m)"
+    proof (induct n)
+      case 0 show ?case
+      proof (rule ccontr)
+        assume "\<not> poly (mulexp 0 [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc 0 *** m)"
+        then have "poly q a = 0"
+          by (simp add: poly_add poly_cmult)
+        with `poly q a \<noteq> 0` show False by simp
+      qed
+    next
+      case (Suc n) show ?case
+        by (rule pexp_Suc [THEN ssubst], rule ccontr)
+          (simp add: poly_mult_left_cancel poly_mult_assoc Suc del: pmult_Cons pexp_Suc)
+    qed
+    ultimately show False by simp
+  qed
+  ultimately show ?thesis by (auto simp add: p)
+qed
 
-lemma poly_order: "poly p \<noteq> poly [] \<Longrightarrow>
-    \<exists>! n. ([-a, (1::'a::{idom,ring_char_0})] %^ n) divides p \<and> \<not> (([-a, 1] %^ Suc n) divides p)"
+lemma (in semiring_1) poly_one_divides[simp]: "[1] divides p"
+  by (auto simp add: divides_def)
+
+lemma (in idom_char_0) poly_order:
+  "poly p \<noteq> poly [] \<Longrightarrow> \<exists>!n. ([-a, 1] %^ n) divides p \<and> \<not> (([-a, 1] %^ Suc n) divides p)"
   apply (auto intro: poly_order_exists simp add: less_linear simp del: pmult_Cons pexp_Suc)
   apply (cut_tac x = y and y = n in less_linear)
   apply (drule_tac m = n in poly_exp_divides)
   apply (auto dest: Suc_le_eq [THEN iffD2, THEN [2] poly_exp_divides]
-    simp del: pmult_Cons pexp_Suc)
+              simp del: pmult_Cons pexp_Suc)
   done
 
 text{*Order*}
 
-lemma some1_equalityD: "n = (SOME n. P n) \<Longrightarrow> EX! n. P n \<Longrightarrow> P n"
+lemma some1_equalityD: "n = (SOME n. P n) \<Longrightarrow> \<exists>!n. P n \<Longrightarrow> P n"
   by (blast intro: someI2)
 
-lemma order:
-  "(([-a, (1::'a::{idom,ring_char_0})] %^ n) divides p &
-    ~(([-a, 1] %^ (Suc n)) divides p)) =
-    ((n = order a p) & ~(poly p = poly []))"
+lemma (in idom_char_0) order:
+      "(([-a, 1] %^ n) divides p \<and>
+        ~(([-a, 1] %^ (Suc n)) divides p)) =
+        ((n = order a p) \<and> ~(poly p = poly []))"
   apply (unfold order_def)
   apply (rule iffI)
   apply (blast dest: poly_divides_zero intro!: some1_equality [symmetric] poly_order)
   apply (blast intro!: poly_order [THEN [2] some1_equalityD])
   done
 
-lemma order2: "poly p \<noteq> poly [] \<Longrightarrow>
-  ([-a, (1::'a::{idom,ring_char_0})] %^ (order a p)) divides p &
-    ~(([-a, 1] %^ (Suc(order a p))) divides p)"
+lemma (in idom_char_0) order2:
+  "poly p \<noteq> poly [] \<Longrightarrow>
+    ([-a, 1] %^ (order a p)) divides p \<and> \<not> (([-a, 1] %^ (Suc (order a p))) divides p)"
   by (simp add: order del: pexp_Suc)
 
-lemma order_unique: "poly p \<noteq> poly [] \<Longrightarrow> ([-a, 1] %^ n) divides p \<Longrightarrow>
-  \<not> (([-a, (1::'a::{idom,ring_char_0})] %^ (Suc n)) divides p) \<Longrightarrow> n = order a p"
+lemma (in idom_char_0) order_unique:
+  "poly p \<noteq> poly [] \<Longrightarrow> ([-a, 1] %^ n) divides p \<Longrightarrow> ~(([-a, 1] %^ (Suc n)) divides p) \<Longrightarrow>
+    n = order a p"
   using order [of a n p] by auto
 
-lemma order_unique_lemma:
-  "(poly p \<noteq> poly [] \<and> ([-a, 1] %^ n) divides p \<and>
-    \<not> (([-a, (1::'a::{idom,ring_char_0})] %^ (Suc n)) divides p)) \<Longrightarrow>
+lemma (in idom_char_0) order_unique_lemma:
+  "poly p \<noteq> poly [] \<and> ([-a, 1] %^ n) divides p \<and> ~(([-a, 1] %^ (Suc n)) divides p) \<Longrightarrow>
     n = order a p"
   by (blast intro: order_unique)
 
-lemma order_poly: "poly p = poly q ==> order a p = order a q"
+lemma (in ring_1) order_poly: "poly p = poly q \<Longrightarrow> order a p = order a q"
   by (auto simp add: fun_eq divides_def poly_mult order_def)
 
-lemma pexp_one [simp]: "p %^ (Suc 0) = p"
-  by (induct p) simp_all
+lemma (in semiring_1) pexp_one[simp]: "p %^ (Suc 0) = p"
+  by (induct "p") auto
+
+lemma (in comm_ring_1) lemma_order_root:
+  "0 < n \<and> [- a, 1] %^ n divides p \<and> ~ [- a, 1] %^ (Suc n) divides p \<Longrightarrow> poly p a = 0"
+  by (induct n arbitrary: a p) (auto simp add: divides_def poly_mult simp del: pmult_Cons)
 
-lemma lemma_order_root:
-  "0 < n & [- a, 1] %^ n divides p & ~ [- a, 1] %^ (Suc n) divides p \<Longrightarrow> poly p a = 0"
-  apply (induct n arbitrary: p a)
-  apply blast
-  apply (auto simp add: divides_def poly_mult simp del: pmult_Cons)
+lemma (in idom_char_0) order_root:
+  "poly p a = 0 \<longleftrightarrow> poly p = poly [] \<or> order a p \<noteq> 0"
+  apply (cases "poly p = poly []")
+  apply auto
+  apply (simp add: poly_linear_divides del: pmult_Cons, safe)
+  apply (drule_tac [!] a = a in order2)
+  apply (rule ccontr)
+  apply (simp add: divides_def poly_mult fun_eq del: pmult_Cons, blast)
+  using neq0_conv
+  apply (blast intro: lemma_order_root)
   done
 
-lemma order_root: "poly p a = (0::'a::{idom,ring_char_0}) \<longleftrightarrow> poly p = poly [] \<or> order a p \<noteq> 0"
-  apply (cases "poly p = poly []")
-  apply auto
-  apply (simp add: poly_linear_divides del: pmult_Cons)
-  apply safe
-  apply (drule_tac [!] a = a in order2)
-  apply (rule ccontr)
-  apply (simp add: divides_def poly_mult fun_eq del: pmult_Cons)
-  apply blast
-  using neq0_conv apply (blast intro: lemma_order_root)
-  done
-
-lemma order_divides: "([-a, 1::'a::{idom,ring_char_0}] %^ n) divides p \<longleftrightarrow>
-    poly p = poly [] \<or> n \<le> order a p"
+lemma (in idom_char_0) order_divides:
+  "([-a, 1] %^ n) divides p \<longleftrightarrow> poly p = poly [] \<or> n \<le> order a p"
   apply (cases "poly p = poly []")
   apply auto
   apply (simp add: divides_def fun_eq poly_mult)
   apply (rule_tac x = "[]" in exI)
-  apply (auto dest!: order2 [where a = a] intro: poly_exp_divides simp del: pexp_Suc)
+  apply (auto dest!: order2 [where a=a] intro: poly_exp_divides simp del: pexp_Suc)
   done
 
-lemma order_decomp:
-  "poly p \<noteq> poly [] \<Longrightarrow>
-    \<exists>q. poly p = poly (([-a, 1] %^ (order a p)) *** q) \<and>
-      \<not> ([-a, 1::'a::{idom,ring_char_0}] divides q)"
+lemma (in idom_char_0) order_decomp:
+  "poly p \<noteq> poly [] \<Longrightarrow> \<exists>q. poly p = poly (([-a, 1] %^ (order a p)) *** q) \<and> ~([-a, 1] divides q)"
   apply (unfold divides_def)
   apply (drule order2 [where a = a])
-  apply (simp add: divides_def del: pexp_Suc pmult_Cons)
-  apply safe
-  apply (rule_tac x = q in exI)
-  apply safe
+  apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe)
+  apply (rule_tac x = q in exI, safe)
   apply (drule_tac x = qa in spec)
   apply (auto simp add: poly_mult fun_eq poly_exp mult_ac simp del: pmult_Cons)
   done
 
 text{*Important composition properties of orders.*}
-
-lemma order_mult: "poly (p *** q) \<noteq> poly [] \<Longrightarrow>
-  order a (p *** q) = order a p + order (a::'a::{idom,ring_char_0}) q"
-  apply (cut_tac a = a and p = "p***q" and n = "order a p + order a q" in order)
+lemma order_mult:
+  "poly (p *** q) \<noteq> poly [] \<Longrightarrow>
+    order a (p *** q) = order a p + order (a::'a::{idom_char_0}) q"
+  apply (cut_tac a = a and p = "p *** q" and n = "order a p + order a q" in order)
   apply (auto simp add: poly_entire simp del: pmult_Cons)
   apply (drule_tac a = a in order2)+
   apply safe
-  apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons)
-  apply safe
+  apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe)
   apply (rule_tac x = "qa *** qaa" in exI)
   apply (simp add: poly_mult mult_ac del: pmult_Cons)
   apply (drule_tac a = a in order_decomp)+
   apply safe
-  apply (subgoal_tac "[-a, 1] divides (qa *** qaa) ")
+  apply (subgoal_tac "[-a,1] divides (qa *** qaa) ")
   apply (simp add: poly_primes del: pmult_Cons)
   apply (auto simp add: divides_def simp del: pmult_Cons)
   apply (rule_tac x = qb in exI)
-  apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) =
-    poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))")
-  apply (drule poly_mult_left_cancel [THEN iffD1])
-  apply force
-  apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) =
-    poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ")
-  apply (drule poly_mult_left_cancel [THEN iffD1])
-  apply force
+  apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) = poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))")
+  apply (drule poly_mult_left_cancel [THEN iffD1], force)
+  apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) = poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ")
+  apply (drule poly_mult_left_cancel [THEN iffD1], force)
   apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
   done
 
-lemma order_root2: "poly p \<noteq> poly [] \<Longrightarrow> poly p a = 0 \<longleftrightarrow> order (a::'a::{idom,ring_char_0}) p \<noteq> 0"
+lemma (in idom_char_0) order_mult:
+  assumes "poly (p *** q) \<noteq> poly []"
+  shows "order a (p *** q) = order a p + order a q"
+  using assms
+  apply (cut_tac a = a and p = "pmult p q" and n = "order a p + order a q" in order)
+  apply (auto simp add: poly_entire simp del: pmult_Cons)
+  apply (drule_tac a = a in order2)+
+  apply safe
+  apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe)
+  apply (rule_tac x = "pmult qa qaa" in exI)
+  apply (simp add: poly_mult mult_ac del: pmult_Cons)
+  apply (drule_tac a = a in order_decomp)+
+  apply safe
+  apply (subgoal_tac "[uminus a, one] divides pmult qa qaa")
+  apply (simp add: poly_primes del: pmult_Cons)
+  apply (auto simp add: divides_def simp del: pmult_Cons)
+  apply (rule_tac x = qb in exI)
+  apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa)) =
+    poly (pmult (pexp [uminus a, one] (?order a p)) (pmult [uminus a, one] qb))")
+  apply (drule poly_mult_left_cancel [THEN iffD1], force)
+  apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a q))
+      (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa))) =
+    poly (pmult (pexp [uminus a, one] (order a q))
+      (pmult (pexp [uminus a, one] (order a p)) (pmult [uminus a, one] qb)))")
+  apply (drule poly_mult_left_cancel [THEN iffD1], force)
+  apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
+  done
+
+lemma (in idom_char_0) order_root2: "poly p \<noteq> poly [] \<Longrightarrow> poly p a = 0 \<longleftrightarrow> order a p \<noteq> 0"
   by (rule order_root [THEN ssubst]) auto
 
-lemma pmult_one [simp]: "[1] *** p = p"
-  by auto
+lemma (in semiring_1) pmult_one[simp]: "[1] *** p = p" by auto
 
-lemma poly_Nil_zero: "poly [] = poly [0]"
+lemma (in semiring_0) poly_Nil_zero: "poly [] = poly [0]"
   by (simp add: fun_eq)
 
-lemma rsquarefree_decomp:
-  "rsquarefree p \<Longrightarrow> poly p a = (0::'a::{idom,ring_char_0}) \<Longrightarrow>
+lemma (in idom_char_0) rsquarefree_decomp:
+  "rsquarefree p \<Longrightarrow> poly p a = 0 \<Longrightarrow>
     \<exists>q. poly p = poly ([-a, 1] *** q) \<and> poly q a \<noteq> 0"
-  apply (simp add: rsquarefree_def)
-  apply safe
+  apply (simp add: rsquarefree_def, safe)
   apply (frule_tac a = a in order_decomp)
   apply (drule_tac x = a in spec)
   apply (drule_tac a = a in order_root2 [symmetric])
   apply (auto simp del: pmult_Cons)
-  apply (rule_tac x = q in exI)
-  apply safe
+  apply (rule_tac x = q in exI, safe)
   apply (simp add: poly_mult fun_eq)
   apply (drule_tac p1 = q in poly_linear_divides [THEN iffD1])
-  apply (simp add: divides_def del: pmult_Cons)
-  apply safe
+  apply (simp add: divides_def del: pmult_Cons, safe)
   apply (drule_tac x = "[]" in spec)
   apply (auto simp add: fun_eq)
   done
@@ -729,72 +824,222 @@
 
 text{*Normalization of a polynomial.*}
 
-lemma poly_normalize [simp]: "poly (pnormalize p) = poly p"
+lemma (in semiring_0) poly_normalize[simp]: "poly (pnormalize p) = poly p"
   by (induct p) (auto simp add: fun_eq)
 
-
 text{*The degree of a polynomial.*}
 
-lemma lemma_degree_zero: "list_all (\<lambda>c. c = 0) p \<longleftrightarrow> pnormalize p = []"
+lemma (in semiring_0) lemma_degree_zero: "list_all (%c. c = 0) p \<longleftrightarrow> pnormalize p = []"
   by (induct p) auto
 
-lemma degree_zero: "poly p = poly ([] :: 'a::{idom,ring_char_0} list) \<Longrightarrow> degree p = 0"
-  apply (cases "pnormalize p = []")
-  apply (simp add: degree_def)
-  apply (auto simp add: poly_zero lemma_degree_zero)
-  done
+lemma (in idom_char_0) degree_zero:
+  assumes "poly p = poly []"
+  shows "degree p = 0"
+  using assms
+  by (cases "pnormalize p = []") (auto simp add: degree_def poly_zero lemma_degree_zero)
 
-lemma pnormalize_sing: "pnormalize [x] = [x] \<longleftrightarrow> x \<noteq> 0"
+lemma (in semiring_0) pnormalize_sing: "(pnormalize [x] = [x]) \<longleftrightarrow> x \<noteq> 0"
+  by simp
+
+lemma (in semiring_0) pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> (pnormalize [x, y] = [x, y])"
   by simp
 
-lemma pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> (pnormalize [x, y] = [x, y])"
-  by simp
-
-lemma pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c # p)"
+lemma (in semiring_0) pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c#p)"
   unfolding pnormal_def by simp
 
-lemma pnormal_tail: "p \<noteq> [] \<Longrightarrow> pnormal (c # p) \<Longrightarrow> pnormal p"
-  unfolding pnormal_def
-  apply (cases "pnormalize p = []")
+lemma (in semiring_0) pnormal_tail: "p\<noteq>[] \<Longrightarrow> pnormal (c#p) \<Longrightarrow> pnormal p"
+  unfolding pnormal_def by(auto split: split_if_asm)
+
+
+lemma (in semiring_0) pnormal_last_nonzero: "pnormal p \<Longrightarrow> last p \<noteq> 0"
+  by (induct p) (simp_all add: pnormal_def split: split_if_asm)
+
+lemma (in semiring_0) pnormal_length: "pnormal p \<Longrightarrow> 0 < length p"
+  unfolding pnormal_def length_greater_0_conv by blast
+
+lemma (in semiring_0) pnormal_last_length: "0 < length p \<Longrightarrow> last p \<noteq> 0 \<Longrightarrow> pnormal p"
+  by (induct p) (auto simp: pnormal_def  split: split_if_asm)
+
+
+lemma (in semiring_0) pnormal_id: "pnormal p \<longleftrightarrow> 0 < length p \<and> last p \<noteq> 0"
+  using pnormal_last_length pnormal_length pnormal_last_nonzero by blast
+
+lemma (in idom_char_0) poly_Cons_eq:
+  "poly (c # cs) = poly (d # ds) \<longleftrightarrow> c = d \<and> poly cs = poly ds"
+  (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume eq: ?lhs
+  hence "\<And>x. poly ((c#cs) +++ -- (d#ds)) x = 0"
+    by (simp only: poly_minus poly_add algebra_simps) (simp add: algebra_simps)
+  hence "poly ((c#cs) +++ -- (d#ds)) = poly []" by(simp add: fun_eq_iff)
+  hence "c = d \<and> list_all (\<lambda>x. x=0) ((cs +++ -- ds))"
+    unfolding poly_zero by (simp add: poly_minus_def algebra_simps)
+  hence "c = d \<and> (\<forall>x. poly (cs +++ -- ds) x = 0)"
+    unfolding poly_zero[symmetric] by simp
+  then show ?rhs by (simp add: poly_minus poly_add algebra_simps fun_eq_iff)
+next
+  assume ?rhs
+  then show ?lhs by(simp add:fun_eq_iff)
+qed
+
+lemma (in idom_char_0) pnormalize_unique: "poly p = poly q \<Longrightarrow> pnormalize p = pnormalize q"
+proof (induct q arbitrary: p)
+  case Nil
+  thus ?case by (simp only: poly_zero lemma_degree_zero) simp
+next
+  case (Cons c cs p)
+  thus ?case
+  proof (induct p)
+    case Nil
+    hence "poly [] = poly (c#cs)" by blast
+    then have "poly (c#cs) = poly [] " by simp
+    thus ?case by (simp only: poly_zero lemma_degree_zero) simp
+  next
+    case (Cons d ds)
+    hence eq: "poly (d # ds) = poly (c # cs)" by blast
+    hence eq': "\<And>x. poly (d # ds) x = poly (c # cs) x" by simp
+    hence "poly (d # ds) 0 = poly (c # cs) 0" by blast
+    hence dc: "d = c" by auto
+    with eq have "poly ds = poly cs"
+      unfolding  poly_Cons_eq by simp
+    with Cons.prems have "pnormalize ds = pnormalize cs" by blast
+    with dc show ?case by simp
+  qed
+qed
+
+lemma (in idom_char_0) degree_unique:
+  assumes pq: "poly p = poly q"
+  shows "degree p = degree q"
+  using pnormalize_unique[OF pq] unfolding degree_def by simp
+
+lemma (in semiring_0) pnormalize_length:
+  "length (pnormalize p) \<le> length p" by (induct p) auto
+
+lemma (in semiring_0) last_linear_mul_lemma:
+  "last ((a %* p) +++ (x#(b %* p))) = (if p = [] then x else b * last p)"
+  apply (induct p arbitrary: a x b)
   apply auto
-  apply (cases "c = 0")
+  apply (subgoal_tac "padd (cmult aa p) (times b a # cmult b p) \<noteq> []")
+  apply simp
+  apply (induct_tac p)
   apply auto
   done
 
-lemma pnormal_last_nonzero: "pnormal p \<Longrightarrow> last p \<noteq> 0"
-  apply (induct p)
-  apply (auto simp add: pnormal_def)
-  apply (case_tac "pnormalize p = []")
-  apply auto
-  apply (case_tac "a = 0")
-  apply auto
-  done
+lemma (in semiring_1) last_linear_mul:
+  assumes p: "p \<noteq> []"
+  shows "last ([a,1] *** p) = last p"
+proof -
+  from p obtain c cs where cs: "p = c#cs" by (cases p) auto
+  from cs have eq: "[a,1] *** p = (a %* (c#cs)) +++ (0#(1 %* (c#cs)))"
+    by (simp add: poly_cmult_distr)
+  show ?thesis using cs
+    unfolding eq last_linear_mul_lemma by simp
+qed
+
+lemma (in semiring_0) pnormalize_eq: "last p \<noteq> 0 \<Longrightarrow> pnormalize p = p"
+  by (induct p) (auto split: split_if_asm)
+
+lemma (in semiring_0) last_pnormalize: "pnormalize p \<noteq> [] \<Longrightarrow> last (pnormalize p) \<noteq> 0"
+  by (induct p) auto
+
+lemma (in semiring_0) pnormal_degree: "last p \<noteq> 0 \<Longrightarrow> degree p = length p - 1"
+  using pnormalize_eq[of p] unfolding degree_def by simp
 
-lemma  pnormal_length: "pnormal p \<Longrightarrow> 0 < length p"
-  unfolding pnormal_def length_greater_0_conv by blast
+lemma (in semiring_0) poly_Nil_ext: "poly [] = (\<lambda>x. 0)"
+  by (rule ext) simp
+
+lemma (in idom_char_0) linear_mul_degree:
+  assumes p: "poly p \<noteq> poly []"
+  shows "degree ([a,1] *** p) = degree p + 1"
+proof -
+  from p have pnz: "pnormalize p \<noteq> []"
+    unfolding poly_zero lemma_degree_zero .
+
+  from last_linear_mul[OF pnz, of a] last_pnormalize[OF pnz]
+  have l0: "last ([a, 1] *** pnormalize p) \<noteq> 0" by simp
+  from last_pnormalize[OF pnz] last_linear_mul[OF pnz, of a]
+    pnormal_degree[OF l0] pnormal_degree[OF last_pnormalize[OF pnz]] pnz
+
+  have th: "degree ([a,1] *** pnormalize p) = degree (pnormalize p) + 1"
+    by simp
+
+  have eqs: "poly ([a,1] *** pnormalize p) = poly ([a,1] *** p)"
+    by (rule ext) (simp add: poly_mult poly_add poly_cmult)
+  from degree_unique[OF eqs] th
+  show ?thesis by (simp add: degree_unique[OF poly_normalize])
+qed
 
-lemma pnormal_last_length: "0 < length p \<Longrightarrow> last p \<noteq> 0 \<Longrightarrow> pnormal p"
-  apply (induct p)
-  apply auto
-  apply (case_tac "p = []")
-  apply auto
-  apply (simp add: pnormal_def)
-  apply (rule pnormal_cons)
-  apply auto
-  done
+lemma (in idom_char_0) linear_pow_mul_degree:
+  "degree([a,1] %^n *** p) = (if poly p = poly [] then 0 else degree p + n)"
+proof (induct n arbitrary: a p)
+  case (0 a p)
+  show ?case
+  proof (cases "poly p = poly []")
+    case True
+    then show ?thesis
+      using degree_unique[OF True] by (simp add: degree_def)
+  next
+    case False
+    then show ?thesis by (auto simp add: poly_Nil_ext)
+  qed
+next
+  case (Suc n a p)
+  have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1] %^ n *** ([a,1] *** p))"
+    apply (rule ext)
+    apply (simp add: poly_mult poly_add poly_cmult)
+    apply (simp add: mult_ac add_ac distrib_left)
+    done
+  note deq = degree_unique[OF eq]
+  show ?case
+  proof (cases "poly p = poly []")
+    case True
+    with eq have eq': "poly ([a,1] %^(Suc n) *** p) = poly []"
+      apply -
+      apply (rule ext)
+      apply (simp add: poly_mult poly_cmult poly_add)
+      done
+    from degree_unique[OF eq'] True show ?thesis
+      by (simp add: degree_def)
+  next
+    case False
+    then have ap: "poly ([a,1] *** p) \<noteq> poly []"
+      using poly_mult_not_eq_poly_Nil unfolding poly_entire by auto
+    have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1]%^n *** ([a,1] *** p))"
+      by (rule ext, simp add: poly_mult poly_add poly_exp poly_cmult algebra_simps)
+    from ap have ap': "(poly ([a,1] *** p) = poly []) = False"
+      by blast
+    have th0: "degree ([a,1]%^n *** ([a,1] *** p)) = degree ([a,1] *** p) + n"
+      apply (simp only: Suc.hyps[of a "pmult [a,one] p"] ap')
+      apply simp
+      done
+    from degree_unique[OF eq] ap False th0 linear_mul_degree[OF False, of a]
+    show ?thesis by (auto simp del: poly.simps)
+  qed
+qed
 
-lemma pnormal_id: "pnormal p \<longleftrightarrow> 0 < length p \<and> last p \<noteq> 0"
-  using pnormal_last_length pnormal_length pnormal_last_nonzero by blast
+lemma (in idom_char_0) order_degree:
+  assumes p0: "poly p \<noteq> poly []"
+  shows "order a p \<le> degree p"
+proof -
+  from order2[OF p0, unfolded divides_def]
+  obtain q where q: "poly p = poly ([- a, 1]%^ (order a p) *** q)" by blast
+  {
+    assume "poly q = poly []"
+    with q p0 have False by (simp add: poly_mult poly_entire)
+  }
+  with degree_unique[OF q, unfolded linear_pow_mul_degree] show ?thesis
+    by auto
+qed
 
 text{*Tidier versions of finiteness of roots.*}
 
-lemma poly_roots_finite_set:
-  "poly p \<noteq> poly [] \<Longrightarrow> finite {x::'a::{idom,ring_char_0}. poly p x = 0}"
+lemma (in idom_char_0) poly_roots_finite_set:
+  "poly p \<noteq> poly [] \<Longrightarrow> finite {x. poly p x = 0}"
   unfolding poly_roots_finite .
 
 text{*bound for polynomial.*}
 
-lemma poly_mono: "abs x \<le> k \<Longrightarrow> abs (poly p (x::'a::{linordered_idom})) \<le> poly (map abs p) k"
+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 (rule_tac y = "abs a + abs (x * poly p x)" in order_trans)
@@ -802,7 +1047,6 @@
   apply (auto intro!: mult_mono simp add: abs_mult)
   done
 
-lemma poly_Sing: "poly [c] x = c"
-  by simp
+lemma (in semiring_0) poly_Sing: "poly [c] x = c" by simp
 
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Decision_Procs/Rat_Pair.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -0,0 +1,522 @@
+(*  Title:      HOL/Decision_Procs/Rat_Pair.thy
+    Author:     Amine Chaieb
+*)
+
+header {* Rational numbers as pairs *}
+
+theory Rat_Pair
+imports Complex_Main
+begin
+
+type_synonym Num = "int \<times> int"
+
+abbreviation Num0_syn :: Num  ("0\<^sub>N")
+  where "0\<^sub>N \<equiv> (0, 0)"
+
+abbreviation Numi_syn :: "int \<Rightarrow> Num"  ("'((_)')\<^sub>N")
+  where "(i)\<^sub>N \<equiv> (i, 1)"
+
+definition isnormNum :: "Num \<Rightarrow> bool" where
+  "isnormNum = (\<lambda>(a,b). (if a = 0 then b = 0 else b > 0 \<and> gcd a b = 1))"
+
+definition normNum :: "Num \<Rightarrow> Num" where
+  "normNum = (\<lambda>(a,b).
+    (if a=0 \<or> b = 0 then (0,0) else
+      (let g = gcd a b
+       in if b > 0 then (a div g, b div g) else (- (a div g), - (b div g)))))"
+
+declare gcd_dvd1_int[presburger] gcd_dvd2_int[presburger]
+
+lemma normNum_isnormNum [simp]: "isnormNum (normNum x)"
+proof -
+  obtain a b where x: "x = (a, b)" by (cases x)
+  { assume "a=0 \<or> b = 0" hence ?thesis by (simp add: x normNum_def isnormNum_def) }
+  moreover
+  { assume anz: "a \<noteq> 0" and bnz: "b \<noteq> 0"
+    let ?g = "gcd a b"
+    let ?a' = "a div ?g"
+    let ?b' = "b div ?g"
+    let ?g' = "gcd ?a' ?b'"
+    from anz bnz have "?g \<noteq> 0" by simp  with gcd_ge_0_int[of a b]
+    have gpos: "?g > 0" by arith
+    have gdvd: "?g dvd a" "?g dvd b" by arith+
+    from dvd_mult_div_cancel[OF gdvd(1)] dvd_mult_div_cancel[OF gdvd(2)] anz bnz
+    have nz': "?a' \<noteq> 0" "?b' \<noteq> 0" by - (rule notI, simp)+
+    from anz bnz have stupid: "a \<noteq> 0 \<or> b \<noteq> 0" by arith
+    from div_gcd_coprime_int[OF stupid] have gp1: "?g' = 1" .
+    from bnz have "b < 0 \<or> b > 0" by arith
+    moreover
+    { assume b: "b > 0"
+      from b have "?b' \<ge> 0"
+        by (presburger add: pos_imp_zdiv_nonneg_iff[OF gpos])
+      with nz' have b': "?b' > 0" by arith
+      from b b' anz bnz nz' gp1 have ?thesis
+        by (simp add: x isnormNum_def normNum_def Let_def split_def) }
+    moreover {
+      assume b: "b < 0"
+      { assume b': "?b' \<ge> 0"
+        from gpos have th: "?g \<ge> 0" by arith
+        from mult_nonneg_nonneg[OF th b'] dvd_mult_div_cancel[OF gdvd(2)]
+        have False using b by arith }
+      hence b': "?b' < 0" by (presburger add: linorder_not_le[symmetric])
+      from anz bnz nz' b b' gp1 have ?thesis
+        by (simp add: x isnormNum_def normNum_def Let_def split_def) }
+    ultimately have ?thesis by blast
+  }
+  ultimately show ?thesis by blast
+qed
+
+text {* Arithmetic over Num *}
+
+definition Nadd :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "+\<^sub>N" 60) where
+  "Nadd = (\<lambda>(a,b) (a',b'). if a = 0 \<or> b = 0 then normNum(a',b')
+    else if a'=0 \<or> b' = 0 then normNum(a,b)
+    else normNum(a*b' + b*a', b*b'))"
+
+definition Nmul :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "*\<^sub>N" 60) where
+  "Nmul = (\<lambda>(a,b) (a',b'). let g = gcd (a*a') (b*b')
+    in (a*a' div g, b*b' div g))"
+
+definition Nneg :: "Num \<Rightarrow> Num" ("~\<^sub>N")
+  where "Nneg \<equiv> (\<lambda>(a,b). (-a,b))"
+
+definition Nsub :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "-\<^sub>N" 60)
+  where "Nsub = (\<lambda>a b. a +\<^sub>N ~\<^sub>N b)"
+
+definition Ninv :: "Num \<Rightarrow> Num"
+  where "Ninv = (\<lambda>(a,b). if a < 0 then (-b, \<bar>a\<bar>) else (b,a))"
+
+definition Ndiv :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "\<div>\<^sub>N" 60)
+  where "Ndiv = (\<lambda>a b. a *\<^sub>N Ninv b)"
+
+lemma Nneg_normN[simp]: "isnormNum x \<Longrightarrow> isnormNum (~\<^sub>N x)"
+  by (simp add: isnormNum_def Nneg_def split_def)
+
+lemma Nadd_normN[simp]: "isnormNum (x +\<^sub>N y)"
+  by (simp add: Nadd_def split_def)
+
+lemma Nsub_normN[simp]: "\<lbrakk> isnormNum y\<rbrakk> \<Longrightarrow> isnormNum (x -\<^sub>N y)"
+  by (simp add: Nsub_def split_def)
+
+lemma Nmul_normN[simp]:
+  assumes xn: "isnormNum x" and yn: "isnormNum y"
+  shows "isnormNum (x *\<^sub>N y)"
+proof -
+  obtain a b where x: "x = (a, b)" by (cases x)
+  obtain a' b' where y: "y = (a', b')" by (cases y)
+  { assume "a = 0"
+    hence ?thesis using xn x y
+      by (simp add: isnormNum_def Let_def Nmul_def split_def) }
+  moreover
+  { assume "a' = 0"
+    hence ?thesis using yn x y
+      by (simp add: isnormNum_def Let_def Nmul_def split_def) }
+  moreover
+  { assume a: "a \<noteq>0" and a': "a'\<noteq>0"
+    hence bp: "b > 0" "b' > 0" using xn yn x y by (simp_all add: isnormNum_def)
+    from mult_pos_pos[OF bp] have "x *\<^sub>N y = normNum (a * a', b * b')"
+      using x y a a' bp by (simp add: Nmul_def Let_def split_def normNum_def)
+    hence ?thesis by simp }
+  ultimately show ?thesis by blast
+qed
+
+lemma Ninv_normN[simp]: "isnormNum x \<Longrightarrow> isnormNum (Ninv x)"
+  by (simp add: Ninv_def isnormNum_def split_def)
+    (cases "fst x = 0", auto simp add: gcd_commute_int)
+
+lemma isnormNum_int[simp]:
+  "isnormNum 0\<^sub>N" "isnormNum ((1::int)\<^sub>N)" "i \<noteq> 0 \<Longrightarrow> isnormNum (i)\<^sub>N"
+  by (simp_all add: isnormNum_def)
+
+
+text {* Relations over Num *}
+
+definition Nlt0:: "Num \<Rightarrow> bool"  ("0>\<^sub>N")
+  where "Nlt0 = (\<lambda>(a,b). a < 0)"
+
+definition Nle0:: "Num \<Rightarrow> bool"  ("0\<ge>\<^sub>N")
+  where "Nle0 = (\<lambda>(a,b). a \<le> 0)"
+
+definition Ngt0:: "Num \<Rightarrow> bool"  ("0<\<^sub>N")
+  where "Ngt0 = (\<lambda>(a,b). a > 0)"
+
+definition Nge0:: "Num \<Rightarrow> bool"  ("0\<le>\<^sub>N")
+  where "Nge0 = (\<lambda>(a,b). a \<ge> 0)"
+
+definition Nlt :: "Num \<Rightarrow> Num \<Rightarrow> bool"  (infix "<\<^sub>N" 55)
+  where "Nlt = (\<lambda>a b. 0>\<^sub>N (a -\<^sub>N b))"
+
+definition Nle :: "Num \<Rightarrow> Num \<Rightarrow> bool"  (infix "\<le>\<^sub>N" 55)
+  where "Nle = (\<lambda>a b. 0\<ge>\<^sub>N (a -\<^sub>N b))"
+
+definition "INum = (\<lambda>(a,b). of_int a / of_int b)"
+
+lemma INum_int [simp]: "INum (i)\<^sub>N = ((of_int i) ::'a::field)" "INum 0\<^sub>N = (0::'a::field)"
+  by (simp_all add: INum_def)
+
+lemma isnormNum_unique[simp]:
+  assumes na: "isnormNum x" and nb: "isnormNum y"
+  shows "((INum x ::'a::{field_char_0, field_inverse_zero}) = INum y) = (x = y)" (is "?lhs = ?rhs")
+proof
+  obtain a b where x: "x = (a, b)" by (cases x)
+  obtain a' b' where y: "y = (a', b')" by (cases y)
+  assume H: ?lhs
+  { assume "a = 0 \<or> b = 0 \<or> a' = 0 \<or> b' = 0"
+    hence ?rhs using na nb H
+      by (simp add: x y INum_def split_def isnormNum_def split: split_if_asm) }
+  moreover
+  { assume az: "a \<noteq> 0" and bz: "b \<noteq> 0" and a'z: "a'\<noteq>0" and b'z: "b'\<noteq>0"
+    from az bz a'z b'z na nb have pos: "b > 0" "b' > 0" by (simp_all add: x y isnormNum_def)
+    from H bz b'z have eq: "a * b' = a'*b"
+      by (simp add: x y INum_def eq_divide_eq divide_eq_eq of_int_mult[symmetric] del: of_int_mult)
+    from az a'z na nb have gcd1: "gcd a b = 1" "gcd b a = 1" "gcd a' b' = 1" "gcd b' a' = 1"
+      by (simp_all add: x y isnormNum_def add: gcd_commute_int)
+    from eq have raw_dvd: "a dvd a' * b" "b dvd b' * a" "a' dvd a * b'" "b' dvd b * a'"
+      apply -
+      apply algebra
+      apply algebra
+      apply simp
+      apply algebra
+      done
+    from zdvd_antisym_abs[OF coprime_dvd_mult_int[OF gcd1(2) raw_dvd(2)]
+        coprime_dvd_mult_int[OF gcd1(4) raw_dvd(4)]]
+      have eq1: "b = b'" using pos by arith
+      with eq have "a = a'" using pos by simp
+      with eq1 have ?rhs by (simp add: x y) }
+  ultimately show ?rhs by blast
+next
+  assume ?rhs thus ?lhs by simp
+qed
+
+
+lemma isnormNum0[simp]:
+    "isnormNum x \<Longrightarrow> (INum x = (0::'a::{field_char_0, field_inverse_zero})) = (x = 0\<^sub>N)"
+  unfolding INum_int(2)[symmetric]
+  by (rule isnormNum_unique) simp_all
+
+lemma of_int_div_aux: "d ~= 0 ==> ((of_int x)::'a::field_char_0) / (of_int d) =
+    of_int (x div d) + (of_int (x mod d)) / ((of_int d)::'a)"
+proof -
+  assume "d ~= 0"
+  let ?t = "of_int (x div d) * ((of_int d)::'a) + of_int(x mod d)"
+  let ?f = "\<lambda>x. x / of_int d"
+  have "x = (x div d) * d + x mod d"
+    by auto
+  then have eq: "of_int x = ?t"
+    by (simp only: of_int_mult[symmetric] of_int_add [symmetric])
+  then have "of_int x / of_int d = ?t / of_int d"
+    using cong[OF refl[of ?f] eq] by simp
+  then show ?thesis by (simp add: add_divide_distrib algebra_simps `d ~= 0`)
+qed
+
+lemma of_int_div: "(d::int) ~= 0 ==> d dvd n ==>
+    (of_int(n div d)::'a::field_char_0) = of_int n / of_int d"
+  apply (frule of_int_div_aux [of d n, where ?'a = 'a])
+  apply simp
+  apply (simp add: dvd_eq_mod_eq_0)
+  done
+
+
+lemma normNum[simp]: "INum (normNum x) = (INum x :: 'a::{field_char_0, field_inverse_zero})"
+proof -
+  obtain a b where x: "x = (a, b)" by (cases x)
+  { assume "a = 0 \<or> b = 0"
+    hence ?thesis by (simp add: x INum_def normNum_def split_def Let_def) }
+  moreover
+  { assume a: "a \<noteq> 0" and b: "b \<noteq> 0"
+    let ?g = "gcd a b"
+    from a b have g: "?g \<noteq> 0"by simp
+    from of_int_div[OF g, where ?'a = 'a]
+    have ?thesis by (auto simp add: x INum_def normNum_def split_def Let_def) }
+  ultimately show ?thesis by blast
+qed
+
+lemma INum_normNum_iff:
+  "(INum x ::'a::{field_char_0, field_inverse_zero}) = INum y \<longleftrightarrow> normNum x = normNum y"
+  (is "?lhs = ?rhs")
+proof -
+  have "normNum x = normNum y \<longleftrightarrow> (INum (normNum x) :: 'a) = INum (normNum y)"
+    by (simp del: normNum)
+  also have "\<dots> = ?lhs" by simp
+  finally show ?thesis by simp
+qed
+
+lemma Nadd[simp]: "INum (x +\<^sub>N y) = INum x + (INum y :: 'a :: {field_char_0, field_inverse_zero})"
+proof -
+  let ?z = "0:: 'a"
+  obtain a b where x: "x = (a, b)" by (cases x)
+  obtain a' b' where y: "y = (a', b')" by (cases y)
+  { assume "a=0 \<or> a'= 0 \<or> b =0 \<or> b' = 0"
+    hence ?thesis
+      apply (cases "a=0", simp_all add: x y Nadd_def)
+      apply (cases "b= 0", simp_all add: INum_def)
+       apply (cases "a'= 0", simp_all)
+       apply (cases "b'= 0", simp_all)
+       done }
+  moreover
+  { assume aa': "a \<noteq> 0" "a'\<noteq> 0" and bb': "b \<noteq> 0" "b' \<noteq> 0"
+    { assume z: "a * b' + b * a' = 0"
+      hence "of_int (a*b' + b*a') / (of_int b* of_int b') = ?z" by simp
+      hence "of_int b' * of_int a / (of_int b * of_int b') +
+          of_int b * of_int a' / (of_int b * of_int b') = ?z"
+        by (simp add:add_divide_distrib)
+      hence th: "of_int a / of_int b + of_int a' / of_int b' = ?z" using bb' aa'
+        by simp
+      from z aa' bb' have ?thesis
+        by (simp add: x y th Nadd_def normNum_def INum_def split_def) }
+    moreover {
+      assume z: "a * b' + b * a' \<noteq> 0"
+      let ?g = "gcd (a * b' + b * a') (b * b')"
+      have gz: "?g \<noteq> 0" using z by simp
+      have ?thesis using aa' bb' z gz
+        of_int_div[where ?'a = 'a, OF gz gcd_dvd1_int[where x="a * b' + b * a'" and y="b*b'"]]
+        of_int_div[where ?'a = 'a, OF gz gcd_dvd2_int[where x="a * b' + b * a'" and y="b*b'"]]
+        by (simp add: x y Nadd_def INum_def normNum_def Let_def) (simp add: field_simps)
+    }
+    ultimately have ?thesis using aa' bb'
+      by (simp add: x y Nadd_def INum_def normNum_def Let_def) }
+  ultimately show ?thesis by blast
+qed
+
+lemma Nmul[simp]: "INum (x *\<^sub>N y) = INum x * (INum y:: 'a :: {field_char_0, field_inverse_zero})"
+proof -
+  let ?z = "0::'a"
+  obtain a b where x: "x = (a, b)" by (cases x)
+  obtain a' b' where y: "y = (a', b')" by (cases y)
+  { assume "a=0 \<or> a'= 0 \<or> b = 0 \<or> b' = 0"
+    hence ?thesis
+      apply (cases "a=0", simp_all add: x y Nmul_def INum_def Let_def)
+      apply (cases "b=0", simp_all)
+      apply (cases "a'=0", simp_all)
+      done }
+  moreover
+  { assume z: "a \<noteq> 0" "a' \<noteq> 0" "b \<noteq> 0" "b' \<noteq> 0"
+    let ?g="gcd (a*a') (b*b')"
+    have gz: "?g \<noteq> 0" using z by simp
+    from z of_int_div[where ?'a = 'a, OF gz gcd_dvd1_int[where x="a*a'" and y="b*b'"]]
+      of_int_div[where ?'a = 'a , OF gz gcd_dvd2_int[where x="a*a'" and y="b*b'"]]
+    have ?thesis by (simp add: Nmul_def x y Let_def INum_def) }
+  ultimately show ?thesis by blast
+qed
+
+lemma Nneg[simp]: "INum (~\<^sub>N x) = - (INum x ::'a:: field)"
+  by (simp add: Nneg_def split_def INum_def)
+
+lemma Nsub[simp]: "INum (x -\<^sub>N y) = INum x - (INum y:: 'a :: {field_char_0, field_inverse_zero})"
+  by (simp add: Nsub_def split_def)
+
+lemma Ninv[simp]: "INum (Ninv x) = (1::'a :: field_inverse_zero) / (INum x)"
+  by (simp add: Ninv_def INum_def split_def)
+
+lemma Ndiv[simp]: "INum (x \<div>\<^sub>N y) = INum x / (INum y ::'a :: {field_char_0, field_inverse_zero})"
+  by (simp add: Ndiv_def)
+
+lemma Nlt0_iff[simp]:
+  assumes nx: "isnormNum x"
+  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero})< 0) = 0>\<^sub>N x"
+proof -
+  obtain a b where x: "x = (a, b)" by (cases x)
+  { assume "a = 0" hence ?thesis by (simp add: x Nlt0_def INum_def) }
+  moreover
+  { assume a: "a \<noteq> 0" hence b: "(of_int b::'a) > 0"
+      using nx by (simp add: x isnormNum_def)
+    from pos_divide_less_eq[OF b, where b="of_int a" and a="0::'a"]
+    have ?thesis by (simp add: x Nlt0_def INum_def) }
+  ultimately show ?thesis by blast
+qed
+
+lemma Nle0_iff[simp]:
+  assumes nx: "isnormNum x"
+  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero}) \<le> 0) = 0\<ge>\<^sub>N x"
+proof -
+  obtain a b where x: "x = (a, b)" by (cases x)
+  { assume "a = 0" hence ?thesis by (simp add: x Nle0_def INum_def) }
+  moreover
+  { assume a: "a \<noteq> 0" hence b: "(of_int b :: 'a) > 0"
+      using nx by (simp add: x isnormNum_def)
+    from pos_divide_le_eq[OF b, where b="of_int a" and a="0::'a"]
+    have ?thesis by (simp add: x Nle0_def INum_def) }
+  ultimately show ?thesis by blast
+qed
+
+lemma Ngt0_iff[simp]:
+  assumes nx: "isnormNum x"
+  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero})> 0) = 0<\<^sub>N x"
+proof -
+  obtain a b where x: "x = (a, b)" by (cases x)
+  { assume "a = 0" hence ?thesis by (simp add: x Ngt0_def INum_def) }
+  moreover
+  { assume a: "a \<noteq> 0" hence b: "(of_int b::'a) > 0" using nx
+      by (simp add: x isnormNum_def)
+    from pos_less_divide_eq[OF b, where b="of_int a" and a="0::'a"]
+    have ?thesis by (simp add: x Ngt0_def INum_def) }
+  ultimately show ?thesis by blast
+qed
+
+lemma Nge0_iff[simp]:
+  assumes nx: "isnormNum x"
+  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero}) \<ge> 0) = 0\<le>\<^sub>N x"
+proof -
+  obtain a b where x: "x = (a, b)" by (cases x)
+  { assume "a = 0" hence ?thesis by (simp add: x Nge0_def INum_def) }
+  moreover
+  { assume "a \<noteq> 0" hence b: "(of_int b::'a) > 0" using nx
+      by (simp add: x isnormNum_def)
+    from pos_le_divide_eq[OF b, where b="of_int a" and a="0::'a"]
+    have ?thesis by (simp add: x Nge0_def INum_def) }
+  ultimately show ?thesis by blast
+qed
+
+lemma Nlt_iff[simp]:
+  assumes nx: "isnormNum x" and ny: "isnormNum y"
+  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero}) < INum y) = (x <\<^sub>N y)"
+proof -
+  let ?z = "0::'a"
+  have "((INum x ::'a) < INum y) = (INum (x -\<^sub>N y) < ?z)"
+    using nx ny by simp
+  also have "\<dots> = (0>\<^sub>N (x -\<^sub>N y))"
+    using Nlt0_iff[OF Nsub_normN[OF ny]] by simp
+  finally show ?thesis by (simp add: Nlt_def)
+qed
+
+lemma Nle_iff[simp]:
+  assumes nx: "isnormNum x" and ny: "isnormNum y"
+  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero})\<le> INum y) = (x \<le>\<^sub>N y)"
+proof -
+  have "((INum x ::'a) \<le> INum y) = (INum (x -\<^sub>N y) \<le> (0::'a))"
+    using nx ny by simp
+  also have "\<dots> = (0\<ge>\<^sub>N (x -\<^sub>N y))"
+    using Nle0_iff[OF Nsub_normN[OF ny]] by simp
+  finally show ?thesis by (simp add: Nle_def)
+qed
+
+lemma Nadd_commute:
+  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
+  shows "x +\<^sub>N y = y +\<^sub>N x"
+proof -
+  have n: "isnormNum (x +\<^sub>N y)" "isnormNum (y +\<^sub>N x)" by simp_all
+  have "(INum (x +\<^sub>N y)::'a) = INum (y +\<^sub>N x)" by simp
+  with isnormNum_unique[OF n] show ?thesis by simp
+qed
+
+lemma [simp]:
+  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
+  shows "(0, b) +\<^sub>N y = normNum y"
+    and "(a, 0) +\<^sub>N y = normNum y"
+    and "x +\<^sub>N (0, b) = normNum x"
+    and "x +\<^sub>N (a, 0) = normNum x"
+  apply (simp add: Nadd_def split_def)
+  apply (simp add: Nadd_def split_def)
+  apply (subst Nadd_commute, simp add: Nadd_def split_def)
+  apply (subst Nadd_commute, simp add: Nadd_def split_def)
+  done
+
+lemma normNum_nilpotent_aux[simp]:
+  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
+  assumes nx: "isnormNum x"
+  shows "normNum x = x"
+proof -
+  let ?a = "normNum x"
+  have n: "isnormNum ?a" by simp
+  have th: "INum ?a = (INum x ::'a)" by simp
+  with isnormNum_unique[OF n nx] show ?thesis by simp
+qed
+
+lemma normNum_nilpotent[simp]:
+  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
+  shows "normNum (normNum x) = normNum x"
+  by simp
+
+lemma normNum0[simp]: "normNum (0,b) = 0\<^sub>N" "normNum (a,0) = 0\<^sub>N"
+  by (simp_all add: normNum_def)
+
+lemma normNum_Nadd:
+  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
+  shows "normNum (x +\<^sub>N y) = x +\<^sub>N y" by simp
+
+lemma Nadd_normNum1[simp]:
+  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
+  shows "normNum x +\<^sub>N y = x +\<^sub>N y"
+proof -
+  have n: "isnormNum (normNum x +\<^sub>N y)" "isnormNum (x +\<^sub>N y)" by simp_all
+  have "INum (normNum x +\<^sub>N y) = INum x + (INum y :: 'a)" by simp
+  also have "\<dots> = INum (x +\<^sub>N y)" by simp
+  finally show ?thesis using isnormNum_unique[OF n] by simp
+qed
+
+lemma Nadd_normNum2[simp]:
+  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
+  shows "x +\<^sub>N normNum y = x +\<^sub>N y"
+proof -
+  have n: "isnormNum (x +\<^sub>N normNum y)" "isnormNum (x +\<^sub>N y)" by simp_all
+  have "INum (x +\<^sub>N normNum y) = INum x + (INum y :: 'a)" by simp
+  also have "\<dots> = INum (x +\<^sub>N y)" by simp
+  finally show ?thesis using isnormNum_unique[OF n] by simp
+qed
+
+lemma Nadd_assoc:
+  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
+  shows "x +\<^sub>N y +\<^sub>N z = x +\<^sub>N (y +\<^sub>N z)"
+proof -
+  have n: "isnormNum (x +\<^sub>N y +\<^sub>N z)" "isnormNum (x +\<^sub>N (y +\<^sub>N z))" by simp_all
+  have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" by simp
+  with isnormNum_unique[OF n] show ?thesis by simp
+qed
+
+lemma Nmul_commute: "isnormNum x \<Longrightarrow> isnormNum y \<Longrightarrow> x *\<^sub>N y = y *\<^sub>N x"
+  by (simp add: Nmul_def split_def Let_def gcd_commute_int mult_commute)
+
+lemma Nmul_assoc:
+  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
+  assumes nx: "isnormNum x" and ny: "isnormNum y" and nz: "isnormNum z"
+  shows "x *\<^sub>N y *\<^sub>N z = x *\<^sub>N (y *\<^sub>N z)"
+proof -
+  from nx ny nz have n: "isnormNum (x *\<^sub>N y *\<^sub>N z)" "isnormNum (x *\<^sub>N (y *\<^sub>N z))"
+    by simp_all
+  have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" by simp
+  with isnormNum_unique[OF n] show ?thesis by simp
+qed
+
+lemma Nsub0:
+  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
+  assumes x: "isnormNum x" and y: "isnormNum y"
+  shows "x -\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = y"
+proof -
+  fix h :: 'a
+  from isnormNum_unique[where 'a = 'a, OF Nsub_normN[OF y], where y="0\<^sub>N"]
+  have "(x -\<^sub>N y = 0\<^sub>N) = (INum (x -\<^sub>N y) = (INum 0\<^sub>N :: 'a)) " by simp
+  also have "\<dots> = (INum x = (INum y :: 'a))" by simp
+  also have "\<dots> = (x = y)" using x y by simp
+  finally show ?thesis .
+qed
+
+lemma Nmul0[simp]: "c *\<^sub>N 0\<^sub>N = 0\<^sub>N" " 0\<^sub>N *\<^sub>N c = 0\<^sub>N"
+  by (simp_all add: Nmul_def Let_def split_def)
+
+lemma Nmul_eq0[simp]:
+  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
+  assumes nx: "isnormNum x" and ny: "isnormNum y"
+  shows "x*\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = 0\<^sub>N \<or> y = 0\<^sub>N"
+proof -
+  fix h :: 'a
+  obtain a b where x: "x = (a, b)" by (cases x)
+  obtain a' b' where y: "y = (a', b')" by (cases y)
+  have n0: "isnormNum 0\<^sub>N" by simp
+  show ?thesis using nx ny
+    apply (simp only: isnormNum_unique[where ?'a = 'a, OF  Nmul_normN[OF nx ny] n0, symmetric]
+      Nmul[where ?'a = 'a])
+    apply (simp add: x y INum_def split_def isnormNum_def split: split_if_asm)
+    done
+qed
+
+lemma Nneg_Nneg[simp]: "~\<^sub>N (~\<^sub>N c) = c"
+  by (simp add: Nneg_def split_def)
+
+lemma Nmul1[simp]:
+    "isnormNum c \<Longrightarrow> (1)\<^sub>N *\<^sub>N c = c"
+    "isnormNum c \<Longrightarrow> c *\<^sub>N (1)\<^sub>N = c"
+  apply (simp_all add: Nmul_def Let_def split_def isnormNum_def)
+  apply (cases "fst c = 0", simp_all, cases c, simp_all)+
+  done
+
+end
--- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -5,7 +5,7 @@
 header {* Implementation and verification of multivariate polynomials *}
 
 theory Reflected_Multivariate_Polynomial
-imports Complex_Main "~~/src/HOL/Library/Abstract_Rat" Polynomial_List
+imports Complex_Main Rat_Pair Polynomial_List
 begin
 
 subsection{* Datatype of polynomial expressions *}
--- a/src/HOL/Decision_Procs/mir_tac.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Decision_Procs/mir_tac.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -34,7 +34,7 @@
              @{thm "divide_zero"}, 
              @{thm "divide_divide_eq_left"}, @{thm "times_divide_eq_right"}, 
              @{thm "times_divide_eq_left"}, @{thm "divide_divide_eq_right"},
-             @{thm "diff_minus"}, @{thm "minus_divide_left"}]
+             @{thm uminus_add_conv_diff [symmetric]}, @{thm "minus_divide_left"}]
 val comp_ths = ths @ comp_arith @ @{thms simp_thms};
 
 
--- a/src/HOL/Deriv.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Deriv.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -98,7 +98,7 @@
 
 lemma FDERIV_diff[simp, FDERIV_intros]:
   "(f has_derivative f') F \<Longrightarrow> (g has_derivative g') F \<Longrightarrow> ((\<lambda>x. f x - g x) has_derivative (\<lambda>x. f' x - g' x)) F"
-  by (simp only: diff_minus FDERIV_add FDERIV_minus)
+  by (simp only: diff_conv_add_uminus FDERIV_add FDERIV_minus)
 
 abbreviation
   -- {* Frechet derivative: D is derivative of function f at x within s *}
@@ -718,13 +718,13 @@
       ((%x. (f(x)-f(a)) / (x-a)) -- a --> D)"
 apply (rule iffI)
 apply (drule_tac k="- a" in LIM_offset)
-apply (simp add: diff_minus)
+apply simp
 apply (drule_tac k="a" in LIM_offset)
 apply (simp add: add_commute)
 done
 
 lemma DERIV_iff2: "(DERIV f x :> D) \<longleftrightarrow> (\<lambda>z. (f z - f x) / (z - x)) --x --> D"
-  by (simp add: deriv_def diff_minus [symmetric] DERIV_LIM_iff)
+  by (simp add: deriv_def DERIV_LIM_iff)
 
 lemma DERIV_cong_ev: "x = y \<Longrightarrow> eventually (\<lambda>x. f x = g x) (nhds x) \<Longrightarrow> u = v \<Longrightarrow>
     DERIV f x :> u \<longleftrightarrow> DERIV g y :> v"
@@ -758,8 +758,7 @@
     let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))"
     show "\<forall>z. f z - f x = ?g z * (z-x)" by simp
     show "isCont ?g x" using der
-      by (simp add: isCont_iff DERIV_iff diff_minus
-               cong: LIM_equal [rule_format])
+      by (simp add: isCont_iff DERIV_iff cong: LIM_equal [rule_format])
     show "?g x = l" by simp
   qed
 next
@@ -787,7 +786,7 @@
 proof -
   from l der [THEN DERIV_D, THEN LIM_D [where r = "l"]]
   have "\<exists>s > 0. (\<forall>z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < l)"
-    by (simp add: diff_minus)
+    by simp
   then obtain s
         where s:   "0 < s"
           and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < l"
@@ -798,8 +797,7 @@
     fix h::real
     assume "0 < h" "h < s"
     with all [of h] show "f x < f (x+h)"
-    proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric]
-    split add: split_if_asm)
+    proof (simp add: abs_if pos_less_divide_eq split add: split_if_asm)
       assume "~ (f (x+h) - f x) / h < l" and h: "0 < h"
       with l
       have "0 < (f (x+h) - f x) / h" by arith
@@ -817,7 +815,7 @@
 proof -
   from l der [THEN DERIV_D, THEN LIM_D [where r = "-l"]]
   have "\<exists>s > 0. (\<forall>z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < -l)"
-    by (simp add: diff_minus)
+    by simp
   then obtain s
         where s:   "0 < s"
           and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < -l"
@@ -828,8 +826,7 @@
     fix h::real
     assume "0 < h" "h < s"
     with all [of "-h"] show "f x < f (x-h)"
-    proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric]
-    split add: split_if_asm)
+    proof (simp add: abs_if pos_less_divide_eq split add: split_if_asm)
       assume " - ((f (x-h) - f x) / h) < l" and h: "0 < h"
       with l
       have "0 < (f (x-h) - f x) / h" by arith
@@ -1131,7 +1128,7 @@
 apply (rule linorder_cases [of a b], auto)
 apply (drule_tac [!] f = f in MVT)
 apply (auto dest: DERIV_isCont DERIV_unique simp add: differentiable_def)
-apply (auto dest: DERIV_unique simp add: ring_distribs diff_minus)
+apply (auto dest: DERIV_unique simp add: ring_distribs)
 done
 
 lemma DERIV_const_ratio_const2:
--- a/src/HOL/Divides.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Divides.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -53,6 +53,16 @@
   shows "(a + b * c) div b = c + a div b"
   using assms div_mult_self1 [of b a c] by (simp add: mult_commute)
 
+lemma div_mult_self3 [simp]:
+  assumes "b \<noteq> 0"
+  shows "(c * b + a) div b = c + a div b"
+  using assms by (simp add: add.commute)
+
+lemma div_mult_self4 [simp]:
+  assumes "b \<noteq> 0"
+  shows "(b * c + a) div b = c + a div b"
+  using assms by (simp add: add.commute)
+
 lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b"
 proof (cases "b = 0")
   case True then show ?thesis by simp
@@ -70,9 +80,18 @@
   then show ?thesis by simp
 qed
 
-lemma mod_mult_self2 [simp]: "(a + b * c) mod b = a mod b"
+lemma mod_mult_self2 [simp]: 
+  "(a + b * c) mod b = a mod b"
   by (simp add: mult_commute [of b])
 
+lemma mod_mult_self3 [simp]:
+  "(c * b + a) mod b = a mod b"
+  by (simp add: add.commute)
+
+lemma mod_mult_self4 [simp]:
+  "(b * c + a) mod b = a mod b"
+  by (simp add: add.commute)
+
 lemma div_mult_self1_is_id [simp]: "b \<noteq> 0 \<Longrightarrow> b * a div b = a"
   using div_mult_self2 [of b 0 a] by simp
 
@@ -420,24 +439,23 @@
 
 text {* Subtraction respects modular equivalence. *}
 
-lemma mod_diff_left_eq: "(a - b) mod c = (a mod c - b) mod c"
-  unfolding diff_minus
-  by (intro mod_add_cong mod_minus_cong) simp_all
-
-lemma mod_diff_right_eq: "(a - b) mod c = (a - b mod c) mod c"
-  unfolding diff_minus
-  by (intro mod_add_cong mod_minus_cong) simp_all
-
-lemma mod_diff_eq: "(a - b) mod c = (a mod c - b mod c) mod c"
-  unfolding diff_minus
-  by (intro mod_add_cong mod_minus_cong) simp_all
+lemma mod_diff_left_eq:
+  "(a - b) mod c = (a mod c - b) mod c"
+  using mod_add_cong [of a c "a mod c" "- b" "- b"] by simp
+
+lemma mod_diff_right_eq:
+  "(a - b) mod c = (a - b mod c) mod c"
+  using mod_add_cong [of a c a "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp
+
+lemma mod_diff_eq:
+  "(a - b) mod c = (a mod c - b mod c) mod c"
+  using mod_add_cong [of a c "a mod c" "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp
 
 lemma mod_diff_cong:
   assumes "a mod c = a' mod c"
   assumes "b mod c = b' mod c"
   shows "(a - b) mod c = (a' - b') mod c"
-  unfolding diff_minus using assms
-  by (intro mod_add_cong mod_minus_cong)
+  using assms mod_add_cong [of a c a' "- b" "- b'"] mod_minus_cong [of b c "b'"] by simp
 
 lemma dvd_neg_div: "y dvd x \<Longrightarrow> -x div y = - (x div y)"
 apply (case_tac "y = 0") apply simp
@@ -477,6 +495,34 @@
 lemma mod_minus1_right [simp]: "a mod (-1) = 0"
   using mod_minus_right [of a 1] by simp
 
+lemma minus_mod_self2 [simp]: 
+  "(a - b) mod b = a mod b"
+  by (simp add: mod_diff_right_eq)
+
+lemma minus_mod_self1 [simp]: 
+  "(b - a) mod b = - a mod b"
+  using mod_add_self2 [of "- a" b] by simp
+
+end
+
+class semiring_div_parity = semiring_div + semiring_numeral +
+  assumes parity: "a mod 2 = 0 \<or> a mod 2 = 1"
+begin
+
+lemma parity_cases [case_names even odd]:
+  assumes "a mod 2 = 0 \<Longrightarrow> P"
+  assumes "a mod 2 = 1 \<Longrightarrow> P"
+  shows P
+  using assms parity by blast
+
+lemma not_mod_2_eq_0_eq_1 [simp]:
+  "a mod 2 \<noteq> 0 \<longleftrightarrow> a mod 2 = 1"
+  by (cases a rule: parity_cases) simp_all
+
+lemma not_mod_2_eq_1_eq_0 [simp]:
+  "a mod 2 \<noteq> 1 \<longleftrightarrow> a mod 2 = 0"
+  by (cases a rule: parity_cases) simp_all
+
 end
 
 
@@ -490,7 +536,6 @@
   and less technical class hierarchy.
 *}
 
-
 class semiring_numeral_div = linordered_semidom + minus + semiring_div +
   assumes diff_invert_add1: "a + b = c \<Longrightarrow> a = c - b"
     and le_add_diff_inverse2: "b \<le> a \<Longrightarrow> a - b + b = a"
@@ -510,18 +555,21 @@
   "a - 0 = a"
   by (rule diff_invert_add1 [symmetric]) simp
 
-lemma parity:
-  "a mod 2 = 0 \<or> a mod 2 = 1"
-proof (rule ccontr)
-  assume "\<not> (a mod 2 = 0 \<or> a mod 2 = 1)"
-  then have "a mod 2 \<noteq> 0" and "a mod 2 \<noteq> 1" by simp_all
-  have "0 < 2" by simp
-  with pos_mod_bound pos_mod_sign have "0 \<le> a mod 2" "a mod 2 < 2" by simp_all
-  with `a mod 2 \<noteq> 0` have "0 < a mod 2" by simp
-  with discrete have "1 \<le> a mod 2" by simp
-  with `a mod 2 \<noteq> 1` have "1 < a mod 2" by simp
-  with discrete have "2 \<le> a mod 2" by simp
-  with `a mod 2 < 2` show False by simp
+subclass semiring_div_parity
+proof
+  fix a
+  show "a mod 2 = 0 \<or> a mod 2 = 1"
+  proof (rule ccontr)
+    assume "\<not> (a mod 2 = 0 \<or> a mod 2 = 1)"
+    then have "a mod 2 \<noteq> 0" and "a mod 2 \<noteq> 1" by simp_all
+    have "0 < 2" by simp
+    with pos_mod_bound pos_mod_sign have "0 \<le> a mod 2" "a mod 2 < 2" by simp_all
+    with `a mod 2 \<noteq> 0` have "0 < a mod 2" by simp
+    with discrete have "1 \<le> a mod 2" by simp
+    with `a mod 2 \<noteq> 1` have "1 < a mod 2" by simp
+    with discrete have "2 \<le> a mod 2" by simp
+    with `a mod 2 < 2` show False by simp
+  qed
 qed
 
 lemma divmod_digit_1:
@@ -1697,7 +1745,7 @@
   val div_mod_eqs = map mk_meta_eq [@{thm div_mod_equality}, @{thm div_mod_equality2}];
 
   val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac 
-    (@{thm diff_minus} :: @{thms add_0s} @ @{thms add_ac}))
+    (@{thm diff_conv_add_uminus} :: @{thms add_0s} @ @{thms add_ac}))
 )
 *}
 
@@ -2568,11 +2616,6 @@
   "Suc 0 mod numeral v' = nat (1 mod numeral v')"
   by (subst nat_mod_distrib) simp_all
 
-lemma mod_2_not_eq_zero_eq_one_int:
-  fixes k :: int
-  shows "k mod 2 \<noteq> 0 \<longleftrightarrow> k mod 2 = 1"
-  by auto
-
 instance int :: semiring_numeral_div
   by intro_classes (auto intro: zmod_le_nonneg_dividend
     simp add: zmult_div_cancel
--- a/src/HOL/Enum.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Enum.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -156,11 +156,11 @@
   "Id = image (\<lambda>x. (x, x)) (set Enum.enum)"
   by (auto intro: imageI in_enum)
 
-lemma tranclp_unfold [code, no_atp]:
+lemma tranclp_unfold [code]:
   "tranclp r a b \<longleftrightarrow> (a, b) \<in> trancl {(x, y). r x y}"
   by (simp add: trancl_def)
 
-lemma rtranclp_rtrancl_eq [code, no_atp]:
+lemma rtranclp_rtrancl_eq [code]:
   "rtranclp r x y \<longleftrightarrow> (x, y) \<in> rtrancl {(x, y). r x y}"
   by (simp add: rtrancl_def)
 
@@ -178,13 +178,9 @@
 
 lemma [code]:
   fixes xs :: "('a::finite \<times> 'a) list"
-  shows "acc (set xs) = bacc (set xs) (card_UNIV TYPE('a))"
+  shows "Wellfounded.acc (set xs) = bacc (set xs) (card_UNIV TYPE('a))"
   by (simp add: card_UNIV_def acc_bacc_eq)
 
-lemma [code]:
-  "accp r = (\<lambda>x. x \<in> acc {(x, y). r x y})"
-  by (simp add: acc_def)
-
 
 subsection {* Default instances for @{class enum} *}
 
--- a/src/HOL/Fields.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Fields.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -152,11 +152,11 @@
 lemma nonzero_minus_divide_divide: "b \<noteq> 0 ==> (-a) / (-b) = a / b"
   by (simp add: divide_inverse nonzero_inverse_minus_eq)
 
-lemma divide_minus_left [simp, no_atp]: "(-a) / b = - (a / b)"
+lemma divide_minus_left [simp]: "(-a) / b = - (a / b)"
   by (simp add: divide_inverse)
 
 lemma diff_divide_distrib: "(a - b) / c = a / c - b / c"
-  by (simp add: diff_minus add_divide_distrib)
+  using add_divide_distrib [of a "- b" c] by simp
 
 lemma nonzero_eq_divide_eq [field_simps]: "c \<noteq> 0 \<Longrightarrow> a = b / c \<longleftrightarrow> a * c = b"
 proof -
@@ -252,7 +252,7 @@
    ==> inverse a + inverse b = (a + b) * inverse a * inverse b"
 by (simp add: division_ring_inverse_add mult_ac)
 
-lemma nonzero_mult_divide_mult_cancel_left [simp, no_atp]:
+lemma nonzero_mult_divide_mult_cancel_left [simp]:
 assumes [simp]: "b\<noteq>0" and [simp]: "c\<noteq>0" shows "(c*a)/(c*b) = a/b"
 proof -
   have "(c*a)/(c*b) = c * a * (inverse b * inverse c)"
@@ -263,7 +263,7 @@
     finally show ?thesis by (simp add: divide_inverse)
 qed
 
-lemma nonzero_mult_divide_mult_cancel_right [simp, no_atp]:
+lemma nonzero_mult_divide_mult_cancel_right [simp]:
   "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (a * c) / (b * c) = a / b"
 by (simp add: mult_commute [of _ c])
 
@@ -275,7 +275,7 @@
   fraction, like a*b*c / x*y*z. The rationale for that is unclear, but
   many proofs seem to need them.*}
 
-lemmas times_divide_eq [no_atp] = times_divide_eq_right times_divide_eq_left
+lemmas times_divide_eq = times_divide_eq_right times_divide_eq_left
 
 lemma add_frac_eq:
   assumes "y \<noteq> 0" and "z \<noteq> 0"
@@ -291,27 +291,27 @@
 
 text{*Special Cancellation Simprules for Division*}
 
-lemma nonzero_mult_divide_cancel_right [simp, no_atp]:
+lemma nonzero_mult_divide_cancel_right [simp]:
   "b \<noteq> 0 \<Longrightarrow> a * b / b = a"
   using nonzero_mult_divide_mult_cancel_right [of 1 b a] by simp
 
-lemma nonzero_mult_divide_cancel_left [simp, no_atp]:
+lemma nonzero_mult_divide_cancel_left [simp]:
   "a \<noteq> 0 \<Longrightarrow> a * b / a = b"
 using nonzero_mult_divide_mult_cancel_left [of 1 a b] by simp
 
-lemma nonzero_divide_mult_cancel_right [simp, no_atp]:
+lemma nonzero_divide_mult_cancel_right [simp]:
   "\<lbrakk>a \<noteq> 0; b \<noteq> 0\<rbrakk> \<Longrightarrow> b / (a * b) = 1 / a"
 using nonzero_mult_divide_mult_cancel_right [of a b 1] by simp
 
-lemma nonzero_divide_mult_cancel_left [simp, no_atp]:
+lemma nonzero_divide_mult_cancel_left [simp]:
   "\<lbrakk>a \<noteq> 0; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / (a * b) = 1 / b"
 using nonzero_mult_divide_mult_cancel_left [of b a 1] by simp
 
-lemma nonzero_mult_divide_mult_cancel_left2 [simp, no_atp]:
+lemma nonzero_mult_divide_mult_cancel_left2 [simp]:
   "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (c * a) / (b * c) = a / b"
 using nonzero_mult_divide_mult_cancel_left [of b c a] by (simp add: mult_ac)
 
-lemma nonzero_mult_divide_mult_cancel_right2 [simp, no_atp]:
+lemma nonzero_mult_divide_mult_cancel_right2 [simp]:
   "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (a * c) / (c * b) = a / b"
 using nonzero_mult_divide_mult_cancel_right [of b c a] by (simp add: mult_ac)
 
@@ -383,18 +383,18 @@
 apply simp_all
 done
 
-lemma divide_divide_eq_right [simp, no_atp]:
+lemma divide_divide_eq_right [simp]:
   "a / (b / c) = (a * c) / b"
   by (simp add: divide_inverse mult_ac)
 
-lemma divide_divide_eq_left [simp, no_atp]:
+lemma divide_divide_eq_left [simp]:
   "(a / b) / c = a / (b * c)"
   by (simp add: divide_inverse mult_assoc)
 
 
 text {*Special Cancellation Simprules for Division*}
 
-lemma mult_divide_mult_cancel_left_if [simp,no_atp]:
+lemma mult_divide_mult_cancel_left_if [simp]:
   shows "(c * a) / (c * b) = (if c = 0 then 0 else a / b)"
   by (simp add: mult_divide_mult_cancel_left)
 
@@ -405,7 +405,7 @@
   "- (a / b) = a / - b"
   by (simp add: divide_inverse)
 
-lemma divide_minus_right [simp, no_atp]:
+lemma divide_minus_right [simp]:
   "a / - b = - (a / b)"
   by (simp add: divide_inverse)
 
@@ -427,29 +427,29 @@
   "inverse x = 1 \<longleftrightarrow> x = 1"
   by (insert inverse_eq_iff_eq [of x 1], simp) 
 
-lemma divide_eq_0_iff [simp, no_atp]:
+lemma divide_eq_0_iff [simp]:
   "a / b = 0 \<longleftrightarrow> a = 0 \<or> b = 0"
   by (simp add: divide_inverse)
 
-lemma divide_cancel_right [simp, no_atp]:
+lemma divide_cancel_right [simp]:
   "a / c = b / c \<longleftrightarrow> c = 0 \<or> a = b"
   apply (cases "c=0", simp)
   apply (simp add: divide_inverse)
   done
 
-lemma divide_cancel_left [simp, no_atp]:
+lemma divide_cancel_left [simp]:
   "c / a = c / b \<longleftrightarrow> c = 0 \<or> a = b" 
   apply (cases "c=0", simp)
   apply (simp add: divide_inverse)
   done
 
-lemma divide_eq_1_iff [simp, no_atp]:
+lemma divide_eq_1_iff [simp]:
   "a / b = 1 \<longleftrightarrow> b \<noteq> 0 \<and> a = b"
   apply (cases "b=0", simp)
   apply (simp add: right_inverse_eq)
   done
 
-lemma one_eq_divide_iff [simp, no_atp]:
+lemma one_eq_divide_iff [simp]:
   "1 = a / b \<longleftrightarrow> b \<noteq> 0 \<and> a = b"
   by (simp add: eq_commute [of 1])
 
@@ -559,7 +559,7 @@
 done
 
 text{*Both premises are essential. Consider -1 and 1.*}
-lemma inverse_less_iff_less [simp,no_atp]:
+lemma inverse_less_iff_less [simp]:
   "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> inverse a < inverse b \<longleftrightarrow> b < a"
   by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less) 
 
@@ -567,7 +567,7 @@
   "a \<le> b \<Longrightarrow> 0 < a \<Longrightarrow> inverse b \<le> inverse a"
   by (force simp add: le_less less_imp_inverse_less)
 
-lemma inverse_le_iff_le [simp,no_atp]:
+lemma inverse_le_iff_le [simp]:
   "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> inverse a \<le> inverse b \<longleftrightarrow> b \<le> a"
   by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le) 
 
@@ -601,7 +601,7 @@
 apply (simp add: nonzero_inverse_minus_eq) 
 done
 
-lemma inverse_less_iff_less_neg [simp,no_atp]:
+lemma inverse_less_iff_less_neg [simp]:
   "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> inverse a < inverse b \<longleftrightarrow> b < a"
 apply (insert inverse_less_iff_less [of "-b" "-a"])
 apply (simp del: inverse_less_iff_less 
@@ -612,7 +612,7 @@
   "a \<le> b \<Longrightarrow> b < 0 ==> inverse b \<le> inverse a"
   by (force simp add: le_less less_imp_inverse_less_neg)
 
-lemma inverse_le_iff_le_neg [simp,no_atp]:
+lemma inverse_le_iff_le_neg [simp]:
   "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> inverse a \<le> inverse b \<longleftrightarrow> b \<le> a"
   by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) 
 
@@ -713,11 +713,9 @@
 sign_simps} to @{text field_simps} because the former can lead to case
 explosions. *}
 
-lemmas sign_simps [no_atp] = algebra_simps
-  zero_less_mult_iff mult_less_0_iff
+lemmas sign_simps = algebra_simps zero_less_mult_iff mult_less_0_iff
 
-lemmas (in -) sign_simps [no_atp] = algebra_simps
-  zero_less_mult_iff mult_less_0_iff
+lemmas (in -) sign_simps = algebra_simps zero_less_mult_iff mult_less_0_iff
 
 (* Only works once linear arithmetic is installed:
 text{*An example:*}
@@ -847,7 +845,7 @@
   fix x y :: 'a
   from less_add_one show "\<exists>y. x < y" .. 
   from less_add_one have "x + (- 1) < (x + 1) + (- 1)" by (rule add_strict_right_mono)
-  then have "x - 1 < x + 1 - 1" by (simp only: diff_minus [symmetric])
+  then have "x - 1 < x + 1 - 1" by simp
   then have "x - 1 < x" by (simp add: algebra_simps)
   then show "\<exists>y. y < x" ..
   show "x < y \<Longrightarrow> \<exists>z>x. z < y" by (blast intro!: less_half_sum gt_half_sum)
@@ -998,13 +996,13 @@
 
 text{*Simplify expressions equated with 1*}
 
-lemma zero_eq_1_divide_iff [simp,no_atp]:
+lemma zero_eq_1_divide_iff [simp]:
      "(0 = 1/a) = (a = 0)"
 apply (cases "a=0", simp)
 apply (auto simp add: nonzero_eq_divide_eq)
 done
 
-lemma one_divide_eq_0_iff [simp,no_atp]:
+lemma one_divide_eq_0_iff [simp]:
      "(1/a = 0) = (a = 0)"
 apply (cases "a=0", simp)
 apply (insert zero_neq_one [THEN not_sym])
@@ -1013,19 +1011,19 @@
 
 text{*Simplify expressions such as @{text "0 < 1/x"} to @{text "0 < x"}*}
 
-lemma zero_le_divide_1_iff [simp, no_atp]:
+lemma zero_le_divide_1_iff [simp]:
   "0 \<le> 1 / a \<longleftrightarrow> 0 \<le> a"
   by (simp add: zero_le_divide_iff)
 
-lemma zero_less_divide_1_iff [simp, no_atp]:
+lemma zero_less_divide_1_iff [simp]:
   "0 < 1 / a \<longleftrightarrow> 0 < a"
   by (simp add: zero_less_divide_iff)
 
-lemma divide_le_0_1_iff [simp, no_atp]:
+lemma divide_le_0_1_iff [simp]:
   "1 / a \<le> 0 \<longleftrightarrow> a \<le> 0"
   by (simp add: divide_le_0_iff)
 
-lemma divide_less_0_1_iff [simp, no_atp]:
+lemma divide_less_0_1_iff [simp]:
   "1 / a < 0 \<longleftrightarrow> a < 0"
   by (simp add: divide_less_0_iff)
 
@@ -1080,62 +1078,62 @@
 
 text{*Simplify quotients that are compared with the value 1.*}
 
-lemma le_divide_eq_1 [no_atp]:
+lemma le_divide_eq_1:
   "(1 \<le> b / a) = ((0 < a & a \<le> b) | (a < 0 & b \<le> a))"
 by (auto simp add: le_divide_eq)
 
-lemma divide_le_eq_1 [no_atp]:
+lemma divide_le_eq_1:
   "(b / a \<le> 1) = ((0 < a & b \<le> a) | (a < 0 & a \<le> b) | a=0)"
 by (auto simp add: divide_le_eq)
 
-lemma less_divide_eq_1 [no_atp]:
+lemma less_divide_eq_1:
   "(1 < b / a) = ((0 < a & a < b) | (a < 0 & b < a))"
 by (auto simp add: less_divide_eq)
 
-lemma divide_less_eq_1 [no_atp]:
+lemma divide_less_eq_1:
   "(b / a < 1) = ((0 < a & b < a) | (a < 0 & a < b) | a=0)"
 by (auto simp add: divide_less_eq)
 
 
 text {*Conditional Simplification Rules: No Case Splits*}
 
-lemma le_divide_eq_1_pos [simp,no_atp]:
+lemma le_divide_eq_1_pos [simp]:
   "0 < a \<Longrightarrow> (1 \<le> b/a) = (a \<le> b)"
 by (auto simp add: le_divide_eq)
 
-lemma le_divide_eq_1_neg [simp,no_atp]:
+lemma le_divide_eq_1_neg [simp]:
   "a < 0 \<Longrightarrow> (1 \<le> b/a) = (b \<le> a)"
 by (auto simp add: le_divide_eq)
 
-lemma divide_le_eq_1_pos [simp,no_atp]:
+lemma divide_le_eq_1_pos [simp]:
   "0 < a \<Longrightarrow> (b/a \<le> 1) = (b \<le> a)"
 by (auto simp add: divide_le_eq)
 
-lemma divide_le_eq_1_neg [simp,no_atp]:
+lemma divide_le_eq_1_neg [simp]:
   "a < 0 \<Longrightarrow> (b/a \<le> 1) = (a \<le> b)"
 by (auto simp add: divide_le_eq)
 
-lemma less_divide_eq_1_pos [simp,no_atp]:
+lemma less_divide_eq_1_pos [simp]:
   "0 < a \<Longrightarrow> (1 < b/a) = (a < b)"
 by (auto simp add: less_divide_eq)
 
-lemma less_divide_eq_1_neg [simp,no_atp]:
+lemma less_divide_eq_1_neg [simp]:
   "a < 0 \<Longrightarrow> (1 < b/a) = (b < a)"
 by (auto simp add: less_divide_eq)
 
-lemma divide_less_eq_1_pos [simp,no_atp]:
+lemma divide_less_eq_1_pos [simp]:
   "0 < a \<Longrightarrow> (b/a < 1) = (b < a)"
 by (auto simp add: divide_less_eq)
 
-lemma divide_less_eq_1_neg [simp,no_atp]:
+lemma divide_less_eq_1_neg [simp]:
   "a < 0 \<Longrightarrow> b/a < 1 <-> a < b"
 by (auto simp add: divide_less_eq)
 
-lemma eq_divide_eq_1 [simp,no_atp]:
+lemma eq_divide_eq_1 [simp]:
   "(1 = b/a) = ((a \<noteq> 0 & a = b))"
 by (auto simp add: eq_divide_eq)
 
-lemma divide_eq_eq_1 [simp,no_atp]:
+lemma divide_eq_eq_1 [simp]:
   "(b/a = 1) = ((a \<noteq> 0 & a = b))"
 by (auto simp add: divide_eq_eq)
 
--- a/src/HOL/Finite_Set.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Finite_Set.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1188,7 +1188,7 @@
   "card A > 0 \<Longrightarrow> finite A"
   by (rule ccontr) simp
 
-lemma card_0_eq [simp, no_atp]:
+lemma card_0_eq [simp]:
   "finite A \<Longrightarrow> card A = 0 \<longleftrightarrow> A = {}"
   by (auto dest: mk_disjoint_insert)
 
@@ -1620,8 +1620,7 @@
   show False by simp (blast dest: Suc_neq_Zero surjD)
 qed
 
-(* Often leads to bogus ATP proofs because of reduced type information, hence no_atp *)
-lemma infinite_UNIV_char_0 [no_atp]:
+lemma infinite_UNIV_char_0:
   "\<not> finite (UNIV :: 'a::semiring_char_0 set)"
 proof
   assume "finite (UNIV :: 'a set)"
--- a/src/HOL/Fun.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Fun.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -775,7 +775,7 @@
 
 subsection {* Cantor's Paradox *}
 
-lemma Cantors_paradox [no_atp]:
+lemma Cantors_paradox:
   "\<not>(\<exists>f. f ` A = Pow A)"
 proof clarify
   fix f assume "f ` A = Pow A" hence *: "Pow A \<le> f ` A" by blast
--- a/src/HOL/GCD.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/GCD.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1555,8 +1555,8 @@
 interpretation gcd_lcm_complete_lattice_nat:
   complete_lattice Gcd Lcm gcd Rings.dvd "\<lambda>m n. m dvd n \<and> \<not> n dvd m" lcm 1 "0::nat"
 where
-  "complete_lattice.INFI Gcd A f = Gcd (f ` A :: nat set)"
-  and "complete_lattice.SUPR Lcm A f = Lcm (f ` A)"
+  "Inf.INFI Gcd A f = Gcd (f ` A :: nat set)"
+  and "Sup.SUPR Lcm A f = Lcm (f ` A)"
 proof -
   show "class.complete_lattice Gcd Lcm gcd Rings.dvd (\<lambda>m n. m dvd n \<and> \<not> n dvd m) lcm 1 (0::nat)"
   proof
@@ -1574,8 +1574,8 @@
   qed
   then interpret gcd_lcm_complete_lattice_nat:
     complete_lattice Gcd Lcm gcd Rings.dvd "\<lambda>m n. m dvd n \<and> \<not> n dvd m" lcm 1 "0::nat" .
-  from gcd_lcm_complete_lattice_nat.INF_def show "complete_lattice.INFI Gcd A f = Gcd (f ` A)" .
-  from gcd_lcm_complete_lattice_nat.SUP_def show "complete_lattice.SUPR Lcm A f = Lcm (f ` A)" .
+  from gcd_lcm_complete_lattice_nat.INF_def show "Inf.INFI Gcd A f = Gcd (f ` A)" .
+  from gcd_lcm_complete_lattice_nat.SUP_def show "Sup.SUPR Lcm A f = Lcm (f ` A)" .
 qed
 
 lemma Lcm_empty_nat: "Lcm {} = (1::nat)"
--- a/src/HOL/Groebner_Basis.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Groebner_Basis.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -10,20 +10,23 @@
 
 subsection {* Groebner Bases *}
 
-lemmas bool_simps = simp_thms(1-34)
+lemmas bool_simps = simp_thms(1-34) -- {* FIXME move to @{theory HOL} *}
+
+lemma nnf_simps: -- {* FIXME shadows fact binding in @{theory HOL} *}
+  "(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)"
+  "(P \<longrightarrow> Q) = (\<not>P \<or> Q)"
+  "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"
+  by blast+
 
 lemma dnf:
-    "(P & (Q | R)) = ((P&Q) | (P&R))" "((Q | R) & P) = ((Q&P) | (R&P))"
-    "(P \<and> Q) = (Q \<and> P)" "(P \<or> Q) = (Q \<or> P)"
+  "(P & (Q | R)) = ((P&Q) | (P&R))"
+  "((Q | R) & P) = ((Q&P) | (R&P))"
+  "(P \<and> Q) = (Q \<and> P)"
+  "(P \<or> Q) = (Q \<or> P)"
   by blast+
 
 lemmas weak_dnf_simps = dnf bool_simps
 
-lemma nnf_simps:
-    "(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)" "(P \<longrightarrow> Q) = (\<not>P \<or> Q)"
-    "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"
-  by blast+
-
 lemma PFalse:
     "P \<equiv> False \<Longrightarrow> \<not> P"
     "\<not> P \<Longrightarrow> (P \<equiv> False)"
--- a/src/HOL/Groups.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Groups.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -321,9 +321,13 @@
 
 class group_add = minus + uminus + monoid_add +
   assumes left_minus [simp]: "- a + a = 0"
-  assumes diff_minus: "a - b = a + (- b)"
+  assumes add_uminus_conv_diff [simp]: "a + (- b) = a - b"
 begin
 
+lemma diff_conv_add_uminus:
+  "a - b = a + (- b)"
+  by simp
+
 lemma minus_unique:
   assumes "a + b = 0" shows "- a = b"
 proof -
@@ -332,8 +336,6 @@
   finally show ?thesis .
 qed
 
-lemmas equals_zero_I = minus_unique (* legacy name *)
-
 lemma minus_zero [simp]: "- 0 = 0"
 proof -
   have "0 + 0 = 0" by (rule add_0_right)
@@ -346,13 +348,17 @@
   thus "- (- a) = a" by (rule minus_unique)
 qed
 
-lemma right_minus [simp]: "a + - a = 0"
+lemma right_minus: "a + - a = 0"
 proof -
   have "a + - a = - (- a) + - a" by simp
   also have "\<dots> = 0" by (rule left_minus)
   finally show ?thesis .
 qed
 
+lemma diff_self [simp]:
+  "a - a = 0"
+  using right_minus [of a] by simp
+
 subclass cancel_semigroup_add
 proof
   fix a b c :: 'a
@@ -367,41 +373,57 @@
   then show "b = c" unfolding add_assoc by simp
 qed
 
-lemma minus_add_cancel: "- a + (a + b) = b"
-by (simp add: add_assoc [symmetric])
+lemma minus_add_cancel [simp]:
+  "- a + (a + b) = b"
+  by (simp add: add_assoc [symmetric])
+
+lemma add_minus_cancel [simp]:
+  "a + (- a + b) = b"
+  by (simp add: add_assoc [symmetric])
 
-lemma add_minus_cancel: "a + (- a + b) = b"
-by (simp add: add_assoc [symmetric])
+lemma diff_add_cancel [simp]:
+  "a - b + b = a"
+  by (simp only: diff_conv_add_uminus add_assoc) simp
 
-lemma minus_add: "- (a + b) = - b + - a"
+lemma add_diff_cancel [simp]:
+  "a + b - b = a"
+  by (simp only: diff_conv_add_uminus add_assoc) simp
+
+lemma minus_add:
+  "- (a + b) = - b + - a"
 proof -
   have "(a + b) + (- b + - a) = 0"
-    by (simp add: add_assoc add_minus_cancel)
-  thus "- (a + b) = - b + - a"
+    by (simp only: add_assoc add_minus_cancel) simp
+  then show "- (a + b) = - b + - a"
     by (rule minus_unique)
 qed
 
-lemma right_minus_eq: "a - b = 0 \<longleftrightarrow> a = b"
+lemma right_minus_eq [simp]:
+  "a - b = 0 \<longleftrightarrow> a = b"
 proof
   assume "a - b = 0"
-  have "a = (a - b) + b" by (simp add:diff_minus add_assoc)
+  have "a = (a - b) + b" by (simp add: add_assoc)
   also have "\<dots> = b" using `a - b = 0` by simp
   finally show "a = b" .
 next
-  assume "a = b" thus "a - b = 0" by (simp add: diff_minus)
+  assume "a = b" thus "a - b = 0" by simp
 qed
 
-lemma diff_self [simp]: "a - a = 0"
-by (simp add: diff_minus)
+lemma eq_iff_diff_eq_0:
+  "a = b \<longleftrightarrow> a - b = 0"
+  by (fact right_minus_eq [symmetric])
 
-lemma diff_0 [simp]: "0 - a = - a"
-by (simp add: diff_minus)
+lemma diff_0 [simp]:
+  "0 - a = - a"
+  by (simp only: diff_conv_add_uminus add_0_left)
 
-lemma diff_0_right [simp]: "a - 0 = a" 
-by (simp add: diff_minus)
+lemma diff_0_right [simp]:
+  "a - 0 = a" 
+  by (simp only: diff_conv_add_uminus minus_zero add_0_right)
 
-lemma diff_minus_eq_add [simp]: "a - - b = a + b"
-by (simp add: diff_minus)
+lemma diff_minus_eq_add [simp]:
+  "a - - b = a + b"
+  by (simp only: diff_conv_add_uminus minus_minus)
 
 lemma neg_equal_iff_equal [simp]:
   "- a = - b \<longleftrightarrow> a = b" 
@@ -416,11 +438,11 @@
 
 lemma neg_equal_0_iff_equal [simp]:
   "- a = 0 \<longleftrightarrow> a = 0"
-by (subst neg_equal_iff_equal [symmetric], simp)
+  by (subst neg_equal_iff_equal [symmetric]) simp
 
 lemma neg_0_equal_iff_equal [simp]:
   "0 = - a \<longleftrightarrow> 0 = a"
-by (subst neg_equal_iff_equal [symmetric], simp)
+  by (subst neg_equal_iff_equal [symmetric]) simp
 
 text{*The next two equations can make the simplifier loop!*}
 
@@ -438,15 +460,8 @@
   thus ?thesis by (simp add: eq_commute)
 qed
 
-lemma diff_add_cancel: "a - b + b = a"
-by (simp add: diff_minus add_assoc)
-
-lemma add_diff_cancel: "a + b - b = a"
-by (simp add: diff_minus add_assoc)
-
-declare diff_minus[symmetric, algebra_simps, field_simps]
-
-lemma eq_neg_iff_add_eq_0: "a = - b \<longleftrightarrow> a + b = 0"
+lemma eq_neg_iff_add_eq_0:
+  "a = - b \<longleftrightarrow> a + b = 0"
 proof
   assume "a = - b" then show "a + b = 0" by simp
 next
@@ -456,72 +471,88 @@
   ultimately show "a = - b" by simp
 qed
 
-lemma add_eq_0_iff: "x + y = 0 \<longleftrightarrow> y = - x"
-  unfolding eq_neg_iff_add_eq_0 [symmetric]
-  by (rule equation_minus_iff)
+lemma add_eq_0_iff2:
+  "a + b = 0 \<longleftrightarrow> a = - b"
+  by (fact eq_neg_iff_add_eq_0 [symmetric])
+
+lemma neg_eq_iff_add_eq_0:
+  "- a = b \<longleftrightarrow> a + b = 0"
+  by (auto simp add: add_eq_0_iff2)
 
-lemma minus_diff_eq [simp]: "- (a - b) = b - a"
-  by (simp add: diff_minus minus_add)
+lemma add_eq_0_iff:
+  "a + b = 0 \<longleftrightarrow> b = - a"
+  by (auto simp add: neg_eq_iff_add_eq_0 [symmetric])
 
-lemma add_diff_eq[algebra_simps, field_simps]: "a + (b - c) = (a + b) - c"
-  by (simp add: diff_minus add_assoc)
+lemma minus_diff_eq [simp]:
+  "- (a - b) = b - a"
+  by (simp only: neg_eq_iff_add_eq_0 diff_conv_add_uminus add_assoc minus_add_cancel) simp
 
-lemma diff_eq_eq[algebra_simps, field_simps]: "a - b = c \<longleftrightarrow> a = c + b"
-  by (auto simp add: diff_minus add_assoc)
+lemma add_diff_eq [algebra_simps, field_simps]:
+  "a + (b - c) = (a + b) - c"
+  by (simp only: diff_conv_add_uminus add_assoc)
 
-lemma eq_diff_eq[algebra_simps, field_simps]: "a = c - b \<longleftrightarrow> a + b = c"
-  by (auto simp add: diff_minus add_assoc)
+lemma diff_add_eq_diff_diff_swap:
+  "a - (b + c) = a - c - b"
+  by (simp only: diff_conv_add_uminus add_assoc minus_add)
 
-lemma diff_diff_eq2[algebra_simps, field_simps]: "a - (b - c) = (a + c) - b"
-  by (simp add: diff_minus minus_add add_assoc)
+lemma diff_eq_eq [algebra_simps, field_simps]:
+  "a - b = c \<longleftrightarrow> a = c + b"
+  by auto
 
-lemma eq_iff_diff_eq_0: "a = b \<longleftrightarrow> a - b = 0"
-  by (fact right_minus_eq [symmetric])
+lemma eq_diff_eq [algebra_simps, field_simps]:
+  "a = c - b \<longleftrightarrow> a + b = c"
+  by auto
+
+lemma diff_diff_eq2 [algebra_simps, field_simps]:
+  "a - (b - c) = (a + c) - b"
+  by (simp only: diff_conv_add_uminus add_assoc) simp
 
 lemma diff_eq_diff_eq:
   "a - b = c - d \<Longrightarrow> a = b \<longleftrightarrow> c = d"
-  by (simp add: eq_iff_diff_eq_0 [of a b] eq_iff_diff_eq_0 [of c d])
+  by (simp only: eq_iff_diff_eq_0 [of a b] eq_iff_diff_eq_0 [of c d])
 
 end
 
 class ab_group_add = minus + uminus + comm_monoid_add +
   assumes ab_left_minus: "- a + a = 0"
-  assumes ab_diff_minus: "a - b = a + (- b)"
+  assumes ab_add_uminus_conv_diff: "a - b = a + (- b)"
 begin
 
 subclass group_add
-  proof qed (simp_all add: ab_left_minus ab_diff_minus)
+  proof qed (simp_all add: ab_left_minus ab_add_uminus_conv_diff)
 
 subclass cancel_comm_monoid_add
 proof
   fix a b c :: 'a
   assume "a + b = a + c"
   then have "- a + a + b = - a + a + c"
-    unfolding add_assoc by simp
+    by (simp only: add_assoc)
   then show "b = c" by simp
 qed
 
-lemma uminus_add_conv_diff[algebra_simps, field_simps]:
+lemma uminus_add_conv_diff [simp]:
   "- a + b = b - a"
-by (simp add:diff_minus add_commute)
+  by (simp add: add_commute)
 
 lemma minus_add_distrib [simp]:
   "- (a + b) = - a + - b"
-by (rule minus_unique) (simp add: add_ac)
+  by (simp add: algebra_simps)
 
-lemma diff_add_eq[algebra_simps, field_simps]: "(a - b) + c = (a + c) - b"
-by (simp add: diff_minus add_ac)
-
-lemma diff_diff_eq[algebra_simps, field_simps]: "(a - b) - c = a - (b + c)"
-by (simp add: diff_minus add_ac)
+lemma diff_add_eq [algebra_simps, field_simps]:
+  "(a - b) + c = (a + c) - b"
+  by (simp add: algebra_simps)
 
-(* FIXME: duplicates right_minus_eq from class group_add *)
-(* but only this one is declared as a simp rule. *)
-lemma diff_eq_0_iff_eq [simp, no_atp]: "a - b = 0 \<longleftrightarrow> a = b"
-  by (rule right_minus_eq)
+lemma diff_diff_eq [algebra_simps, field_simps]:
+  "(a - b) - c = a - (b + c)"
+  by (simp add: algebra_simps)
 
-lemma add_diff_cancel_left: "(c + a) - (c + b) = a - b"
-  by (simp add: diff_minus add_ac)
+lemma diff_add_eq_diff_diff:
+  "a - (b + c) = a - b - c"
+  using diff_add_eq_diff_diff_swap [of a c b] by (simp add: add.commute)
+
+lemma add_diff_cancel_left [simp]:
+  "(c + a) - (c + b) = a - b"
+  by (simp add: algebra_simps)
 
 end
 
@@ -622,19 +653,19 @@
 
 lemma add_less_cancel_left [simp]:
   "c + a < c + b \<longleftrightarrow> a < b"
-by (blast intro: add_less_imp_less_left add_strict_left_mono) 
+  by (blast intro: add_less_imp_less_left add_strict_left_mono) 
 
 lemma add_less_cancel_right [simp]:
   "a + c < b + c \<longleftrightarrow> a < b"
-by (blast intro: add_less_imp_less_right add_strict_right_mono)
+  by (blast intro: add_less_imp_less_right add_strict_right_mono)
 
 lemma add_le_cancel_left [simp]:
   "c + a \<le> c + b \<longleftrightarrow> a \<le> b"
-by (auto, drule add_le_imp_le_left, simp_all add: add_left_mono) 
+  by (auto, drule add_le_imp_le_left, simp_all add: add_left_mono) 
 
 lemma add_le_cancel_right [simp]:
   "a + c \<le> b + c \<longleftrightarrow> a \<le> b"
-by (simp add: add_commute [of a c] add_commute [of b c])
+  by (simp add: add_commute [of a c] add_commute [of b c])
 
 lemma add_le_imp_le_right:
   "a + c \<le> b + c \<Longrightarrow> a \<le> b"
@@ -806,6 +837,22 @@
   then show "x + y = 0" by simp
 qed
 
+lemma add_increasing:
+  "0 \<le> a \<Longrightarrow> b \<le> c \<Longrightarrow> b \<le> a + c"
+  by (insert add_mono [of 0 a b c], simp)
+
+lemma add_increasing2:
+  "0 \<le> c \<Longrightarrow> b \<le> a \<Longrightarrow> b \<le> a + c"
+  by (simp add: add_increasing add_commute [of a])
+
+lemma add_strict_increasing:
+  "0 < a \<Longrightarrow> b \<le> c \<Longrightarrow> b < a + c"
+  by (insert add_less_le_mono [of 0 a b c], simp)
+
+lemma add_strict_increasing2:
+  "0 \<le> a \<Longrightarrow> b < c \<Longrightarrow> b < a + c"
+  by (insert add_le_less_mono [of 0 a b c], simp)
+
 end
 
 class ordered_ab_group_add =
@@ -825,21 +872,53 @@
 
 subclass ordered_comm_monoid_add ..
 
+lemma add_less_same_cancel1 [simp]:
+  "b + a < b \<longleftrightarrow> a < 0"
+  using add_less_cancel_left [of _ _ 0] by simp
+
+lemma add_less_same_cancel2 [simp]:
+  "a + b < b \<longleftrightarrow> a < 0"
+  using add_less_cancel_right [of _ _ 0] by simp
+
+lemma less_add_same_cancel1 [simp]:
+  "a < a + b \<longleftrightarrow> 0 < b"
+  using add_less_cancel_left [of _ 0] by simp
+
+lemma less_add_same_cancel2 [simp]:
+  "a < b + a \<longleftrightarrow> 0 < b"
+  using add_less_cancel_right [of 0] by simp
+
+lemma add_le_same_cancel1 [simp]:
+  "b + a \<le> b \<longleftrightarrow> a \<le> 0"
+  using add_le_cancel_left [of _ _ 0] by simp
+
+lemma add_le_same_cancel2 [simp]:
+  "a + b \<le> b \<longleftrightarrow> a \<le> 0"
+  using add_le_cancel_right [of _ _ 0] by simp
+
+lemma le_add_same_cancel1 [simp]:
+  "a \<le> a + b \<longleftrightarrow> 0 \<le> b"
+  using add_le_cancel_left [of _ 0] by simp
+
+lemma le_add_same_cancel2 [simp]:
+  "a \<le> b + a \<longleftrightarrow> 0 \<le> b"
+  using add_le_cancel_right [of 0] by simp
+
 lemma max_diff_distrib_left:
   shows "max x y - z = max (x - z) (y - z)"
-by (simp add: diff_minus, rule max_add_distrib_left) 
+  using max_add_distrib_left [of x y "- z"] by simp
 
 lemma min_diff_distrib_left:
   shows "min x y - z = min (x - z) (y - z)"
-by (simp add: diff_minus, rule min_add_distrib_left) 
+  using min_add_distrib_left [of x y "- z"] by simp
 
 lemma le_imp_neg_le:
   assumes "a \<le> b" shows "-b \<le> -a"
 proof -
   have "-a+a \<le> -a+b" using `a \<le> b` by (rule add_left_mono) 
-  hence "0 \<le> -a+b" by simp
-  hence "0 + (-b) \<le> (-a + b) + (-b)" by (rule add_right_mono) 
-  thus ?thesis by (simp add: add_assoc)
+  then have "0 \<le> -a+b" by simp
+  then have "0 + (-b) \<le> (-a + b) + (-b)" by (rule add_right_mono) 
+  then show ?thesis by (simp add: algebra_simps)
 qed
 
 lemma neg_le_iff_le [simp]: "- b \<le> - a \<longleftrightarrow> a \<le> b"
@@ -896,35 +975,37 @@
 lemma minus_le_iff: "- a \<le> b \<longleftrightarrow> - b \<le> a"
 by (auto simp add: le_less minus_less_iff)
 
-lemma diff_less_0_iff_less [simp, no_atp]:
+lemma diff_less_0_iff_less [simp]:
   "a - b < 0 \<longleftrightarrow> a < b"
 proof -
-  have "a - b < 0 \<longleftrightarrow> a + (- b) < b + (- b)" by (simp add: diff_minus)
+  have "a - b < 0 \<longleftrightarrow> a + (- b) < b + (- b)" by simp
   also have "... \<longleftrightarrow> a < b" by (simp only: add_less_cancel_right)
   finally show ?thesis .
 qed
 
 lemmas less_iff_diff_less_0 = diff_less_0_iff_less [symmetric]
 
-lemma diff_less_eq[algebra_simps, field_simps]: "a - b < c \<longleftrightarrow> a < c + b"
+lemma diff_less_eq [algebra_simps, field_simps]:
+  "a - b < c \<longleftrightarrow> a < c + b"
 apply (subst less_iff_diff_less_0 [of a])
 apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst])
-apply (simp add: diff_minus add_ac)
+apply (simp add: algebra_simps)
 done
 
-lemma less_diff_eq[algebra_simps, field_simps]: "a < c - b \<longleftrightarrow> a + b < c"
+lemma less_diff_eq[algebra_simps, field_simps]:
+  "a < c - b \<longleftrightarrow> a + b < c"
 apply (subst less_iff_diff_less_0 [of "a + b"])
 apply (subst less_iff_diff_less_0 [of a])
-apply (simp add: diff_minus add_ac)
+apply (simp add: algebra_simps)
 done
 
 lemma diff_le_eq[algebra_simps, field_simps]: "a - b \<le> c \<longleftrightarrow> a \<le> c + b"
-by (auto simp add: le_less diff_less_eq diff_add_cancel add_diff_cancel)
+by (auto simp add: le_less diff_less_eq )
 
 lemma le_diff_eq[algebra_simps, field_simps]: "a \<le> c - b \<longleftrightarrow> a + b \<le> c"
-by (auto simp add: le_less less_diff_eq diff_add_cancel add_diff_cancel)
+by (auto simp add: le_less less_diff_eq)
 
-lemma diff_le_0_iff_le [simp, no_atp]:
+lemma diff_le_0_iff_le [simp]:
   "a - b \<le> 0 \<longleftrightarrow> a \<le> b"
   by (simp add: algebra_simps)
 
@@ -992,63 +1073,6 @@
 
 subclass linordered_cancel_ab_semigroup_add ..
 
-lemma neg_less_eq_nonneg [simp]:
-  "- a \<le> a \<longleftrightarrow> 0 \<le> a"
-proof
-  assume A: "- a \<le> a" show "0 \<le> a"
-  proof (rule classical)
-    assume "\<not> 0 \<le> a"
-    then have "a < 0" by auto
-    with A have "- a < 0" by (rule le_less_trans)
-    then show ?thesis by auto
-  qed
-next
-  assume A: "0 \<le> a" show "- a \<le> a"
-  proof (rule order_trans)
-    show "- a \<le> 0" using A by (simp add: minus_le_iff)
-  next
-    show "0 \<le> a" using A .
-  qed
-qed
-
-lemma neg_less_nonneg [simp]:
-  "- a < a \<longleftrightarrow> 0 < a"
-proof
-  assume A: "- a < a" show "0 < a"
-  proof (rule classical)
-    assume "\<not> 0 < a"
-    then have "a \<le> 0" by auto
-    with A have "- a < 0" by (rule less_le_trans)
-    then show ?thesis by auto
-  qed
-next
-  assume A: "0 < a" show "- a < a"
-  proof (rule less_trans)
-    show "- a < 0" using A by (simp add: minus_le_iff)
-  next
-    show "0 < a" using A .
-  qed
-qed
-
-lemma less_eq_neg_nonpos [simp]:
-  "a \<le> - a \<longleftrightarrow> a \<le> 0"
-proof
-  assume A: "a \<le> - a" show "a \<le> 0"
-  proof (rule classical)
-    assume "\<not> a \<le> 0"
-    then have "0 < a" by auto
-    then have "0 < - a" using A by (rule less_le_trans)
-    then show ?thesis by auto
-  qed
-next
-  assume A: "a \<le> 0" show "a \<le> - a"
-  proof (rule order_trans)
-    show "0 \<le> - a" using A by (simp add: minus_le_iff)
-  next
-    show "a \<le> 0" using A .
-  qed
-qed
-
 lemma equal_neg_zero [simp]:
   "a = - a \<longleftrightarrow> a = 0"
 proof
@@ -1070,6 +1094,37 @@
   "- a = a \<longleftrightarrow> a = 0"
   by (auto dest: sym)
 
+lemma neg_less_eq_nonneg [simp]:
+  "- a \<le> a \<longleftrightarrow> 0 \<le> a"
+proof
+  assume A: "- a \<le> a" show "0 \<le> a"
+  proof (rule classical)
+    assume "\<not> 0 \<le> a"
+    then have "a < 0" by auto
+    with A have "- a < 0" by (rule le_less_trans)
+    then show ?thesis by auto
+  qed
+next
+  assume A: "0 \<le> a" show "- a \<le> a"
+  proof (rule order_trans)
+    show "- a \<le> 0" using A by (simp add: minus_le_iff)
+  next
+    show "0 \<le> a" using A .
+  qed
+qed
+
+lemma neg_less_pos [simp]:
+  "- a < a \<longleftrightarrow> 0 < a"
+  by (auto simp add: less_le)
+
+lemma less_eq_neg_nonpos [simp]:
+  "a \<le> - a \<longleftrightarrow> a \<le> 0"
+  using neg_less_eq_nonneg [of "- a"] by simp
+
+lemma less_neg_neg [simp]:
+  "a < - a \<longleftrightarrow> a < 0"
+  using neg_less_pos [of "- a"] by simp
+
 lemma double_zero [simp]:
   "a + a = 0 \<longleftrightarrow> a = 0"
 proof
@@ -1088,7 +1143,7 @@
   assume "0 < a + a"
   then have "0 - a < a" by (simp only: diff_less_eq)
   then have "- a < a" by simp
-  then show "0 < a" by (simp only: neg_less_nonneg)
+  then show "0 < a" by simp
 next
   assume "0 < a"
   with this have "0 + 0 < a + a"
@@ -1116,24 +1171,6 @@
   then show ?thesis by simp
 qed
 
-lemma le_minus_self_iff:
-  "a \<le> - a \<longleftrightarrow> a \<le> 0"
-proof -
-  from add_le_cancel_left [of "- a" "a + a" 0]
-  have "a \<le> - a \<longleftrightarrow> a + a \<le> 0" 
-    by (simp add: add_assoc [symmetric])
-  thus ?thesis by simp
-qed
-
-lemma minus_le_self_iff:
-  "- a \<le> a \<longleftrightarrow> 0 \<le> a"
-proof -
-  from add_le_cancel_left [of "- a" 0 "a + a"]
-  have "- a \<le> a \<longleftrightarrow> 0 \<le> a + a" 
-    by (simp add: add_assoc [symmetric])
-  thus ?thesis by simp
-qed
-
 lemma minus_max_eq_min:
   "- max x y = min (-x) (-y)"
   by (auto simp add: max_def min_def)
@@ -1144,27 +1181,6 @@
 
 end
 
-context ordered_comm_monoid_add
-begin
-
-lemma add_increasing:
-  "0 \<le> a \<Longrightarrow> b \<le> c \<Longrightarrow> b \<le> a + c"
-  by (insert add_mono [of 0 a b c], simp)
-
-lemma add_increasing2:
-  "0 \<le> c \<Longrightarrow> b \<le> a \<Longrightarrow> b \<le> a + c"
-  by (simp add: add_increasing add_commute [of a])
-
-lemma add_strict_increasing:
-  "0 < a \<Longrightarrow> b \<le> c \<Longrightarrow> b < a + c"
-  by (insert add_less_le_mono [of 0 a b c], simp)
-
-lemma add_strict_increasing2:
-  "0 \<le> a \<Longrightarrow> b < c \<Longrightarrow> b < a + c"
-  by (insert add_le_less_mono [of 0 a b c], simp)
-
-end
-
 class abs =
   fixes abs :: "'a \<Rightarrow> 'a"
 begin
@@ -1231,7 +1247,7 @@
 lemma abs_zero [simp]: "\<bar>0\<bar> = 0"
 by simp
 
-lemma abs_0_eq [simp, no_atp]: "0 = \<bar>a\<bar> \<longleftrightarrow> a = 0"
+lemma abs_0_eq [simp]: "0 = \<bar>a\<bar> \<longleftrightarrow> a = 0"
 proof -
   have "0 = \<bar>a\<bar> \<longleftrightarrow> \<bar>a\<bar> = 0" by (simp only: eq_ac)
   thus ?thesis by simp
@@ -1299,7 +1315,7 @@
 lemma abs_triangle_ineq2: "\<bar>a\<bar> - \<bar>b\<bar> \<le> \<bar>a - b\<bar>"
 proof -
   have "\<bar>a\<bar> = \<bar>b + (a - b)\<bar>"
-    by (simp add: algebra_simps add_diff_cancel)
+    by (simp add: algebra_simps)
   then have "\<bar>a\<bar> \<le> \<bar>b\<bar> + \<bar>a - b\<bar>"
     by (simp add: abs_triangle_ineq)
   then show ?thesis
@@ -1314,14 +1330,14 @@
 
 lemma abs_triangle_ineq4: "\<bar>a - b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
 proof -
-  have "\<bar>a - b\<bar> = \<bar>a + - b\<bar>" by (subst diff_minus, rule refl)
+  have "\<bar>a - b\<bar> = \<bar>a + - b\<bar>" by (simp add: algebra_simps)
   also have "... \<le> \<bar>a\<bar> + \<bar>- b\<bar>" by (rule abs_triangle_ineq)
   finally show ?thesis by simp
 qed
 
 lemma abs_diff_triangle_ineq: "\<bar>a + b - (c + d)\<bar> \<le> \<bar>a - c\<bar> + \<bar>b - d\<bar>"
 proof -
-  have "\<bar>a + b - (c+d)\<bar> = \<bar>(a-c) + (b-d)\<bar>" by (simp add: diff_minus add_ac)
+  have "\<bar>a + b - (c+d)\<bar> = \<bar>(a-c) + (b-d)\<bar>" by (simp add: algebra_simps)
   also have "... \<le> \<bar>a-c\<bar> + \<bar>b-d\<bar>" by (rule abs_triangle_ineq)
   finally show ?thesis .
 qed
@@ -1341,7 +1357,7 @@
 
 subsection {* Tools setup *}
 
-lemma add_mono_thms_linordered_semiring [no_atp]:
+lemma add_mono_thms_linordered_semiring:
   fixes i j k :: "'a\<Colon>ordered_ab_semigroup_add"
   shows "i \<le> j \<and> k \<le> l \<Longrightarrow> i + k \<le> j + l"
     and "i = j \<and> k \<le> l \<Longrightarrow> i + k \<le> j + l"
@@ -1349,7 +1365,7 @@
     and "i = j \<and> k = l \<Longrightarrow> i + k = j + l"
 by (rule add_mono, clarify+)+
 
-lemma add_mono_thms_linordered_field [no_atp]:
+lemma add_mono_thms_linordered_field:
   fixes i j k :: "'a\<Colon>ordered_cancel_ab_semigroup_add"
   shows "i < j \<and> k = l \<Longrightarrow> i + k < j + l"
     and "i = j \<and> k < l \<Longrightarrow> i + k < j + l"
@@ -1362,10 +1378,5 @@
 code_identifier
   code_module Groups \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
 
-
-text {* Legacy *}
-
-lemmas diff_def = diff_minus
-
 end
 
--- a/src/HOL/Hahn_Banach/Bounds.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Hahn_Banach/Bounds.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -57,25 +57,7 @@
   finally show ?thesis .
 qed
 
-lemma lub_compat: "lub A x = isLub UNIV A x"
-proof -
-  have "isUb UNIV A = (\<lambda>x. A *<= x \<and> x \<in> UNIV)"
-    by (rule ext) (simp only: isUb_def)
-  then show ?thesis
-    by (simp only: lub_def isLub_def leastP_def setge_def setle_def) blast
-qed
-
-lemma real_complete:
-  fixes A :: "real set"
-  assumes nonempty: "\<exists>a. a \<in> A"
-    and ex_upper: "\<exists>y. \<forall>a \<in> A. a \<le> y"
-  shows "\<exists>x. lub A x"
-proof -
-  from ex_upper have "\<exists>y. isUb UNIV A y"
-    unfolding isUb_def setle_def by blast
-  with nonempty have "\<exists>x. isLub UNIV A x"
-    by (rule reals_complete)
-  then show ?thesis by (simp only: lub_compat)
-qed
+lemma real_complete: "\<exists>a::real. a \<in> A \<Longrightarrow> \<exists>y. \<forall>a \<in> A. a \<le> y \<Longrightarrow> \<exists>x. lub A x"
+  by (intro exI[of _ "Sup A"]) (auto intro!: cSup_upper cSup_least simp: lub_def)
 
 end
--- a/src/HOL/Hahn_Banach/Vector_Space.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Hahn_Banach/Vector_Space.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -112,7 +112,7 @@
 proof -
   assume x: "x \<in> V"
   have " (a - b) \<cdot> x = (a + - b) \<cdot> x"
-    by (simp add: diff_minus)
+    by simp
   also from x have "\<dots> = a \<cdot> x + (- b) \<cdot> x"
     by (rule add_mult_distrib2)
   also from x have "\<dots> = a \<cdot> x + - (b \<cdot> x)"
--- a/src/HOL/Hilbert_Choice.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Hilbert_Choice.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -741,7 +741,7 @@
 | "bacc r (Suc n) = (bacc r n \<union> {x. \<forall>y. (y, x) \<in> r \<longrightarrow> y \<in> bacc r n})"
 
 lemma bacc_subseteq_acc:
-  "bacc r n \<subseteq> acc r"
+  "bacc r n \<subseteq> Wellfounded.acc r"
   by (induct n) (auto intro: acc.intros)
 
 lemma bacc_mono:
@@ -761,10 +761,10 @@
 
 lemma acc_subseteq_bacc:
   assumes "finite r"
-  shows "acc r \<subseteq> (\<Union>n. bacc r n)"
+  shows "Wellfounded.acc r \<subseteq> (\<Union>n. bacc r n)"
 proof
   fix x
-  assume "x : acc r"
+  assume "x : Wellfounded.acc r"
   then have "\<exists> n. x : bacc r n"
   proof (induct x arbitrary: rule: acc.induct)
     case (accI x)
@@ -788,7 +788,7 @@
 lemma acc_bacc_eq:
   fixes A :: "('a :: finite \<times> 'a) set"
   assumes "finite A"
-  shows "acc A = bacc A (card (UNIV :: 'a set))"
+  shows "Wellfounded.acc A = bacc A (card (UNIV :: 'a set))"
   using assms by (metis acc_subseteq_bacc bacc_subseteq_acc bacc_upper_bound order_eq_iff)
 
 
--- a/src/HOL/IMP/AExp.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/IMP/AExp.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -33,11 +33,12 @@
   "_State" :: "updbinds => 'a" ("<_>")
 translations
   "_State ms" == "_Update <> ms"
+  "_State (_updbinds b bs)" <= "_Update (_State b) bs"
 
 text {* \noindent
   We can now write a series of updates to the function @{text "\<lambda>x. 0"} compactly:
 *}
-lemma "<a := Suc 0, b := 2> = (<> (a := Suc 0)) (b := 2)"
+lemma "<a := 1, b := 2> = (<> (a := 1)) (b := (2::int))"
   by (rule refl)
 
 value "aval (Plus (V ''x'') (N 5)) <''x'' := 7>"
--- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -135,8 +135,6 @@
 
 subsubsection "Ascending Chain Condition"
 
-hide_const (open) acc
-
 abbreviation "strict r == r \<inter> -(r^-1)"
 abbreviation "acc r == wf((strict r)^-1)"
 
--- a/src/HOL/IMP/Big_Step.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/IMP/Big_Step.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -268,11 +268,9 @@
 subsection "Execution is deterministic"
 
 text {* This proof is automatic. *}
-text_raw{*\snip{BigStepDeterministic}{0}{1}{% *}
+
 theorem big_step_determ: "\<lbrakk> (c,s) \<Rightarrow> t; (c,s) \<Rightarrow> u \<rbrakk> \<Longrightarrow> u = t"
   by (induction arbitrary: u rule: big_step.induct) blast+
-text_raw{*}%endsnip*}
-
 
 text {*
   This is the proof as you might present it in a lecture. The remaining
--- a/src/HOL/IMP/Fold.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/IMP/Fold.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1,5 +1,3 @@
-header "Constant Folding"
-
 theory Fold imports Sem_Equiv Vars begin
 
 subsection "Simple folding of arithmetic expressions"
--- a/src/HOL/IMP/Sem_Equiv.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/IMP/Sem_Equiv.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1,9 +1,11 @@
-header "Semantic Equivalence up to a Condition"
+header "Constant Folding"
 
 theory Sem_Equiv
 imports Big_Step
 begin
 
+subsection "Semantic Equivalence up to a Condition"
+
 type_synonym assn = "state \<Rightarrow> bool"
 
 definition
--- a/src/HOL/IMP/Small_Step.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/IMP/Small_Step.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -4,7 +4,6 @@
 
 subsection "The transition relation"
 
-text_raw{*\snip{SmallStepDef}{0}{2}{% *}
 inductive
   small_step :: "com * state \<Rightarrow> com * state \<Rightarrow> bool" (infix "\<rightarrow>" 55)
 where
@@ -18,7 +17,6 @@
 
 While:   "(WHILE b DO c,s) \<rightarrow>
             (IF b THEN c;; WHILE b DO c ELSE SKIP,s)"
-text_raw{*}%endsnip*}
 
 
 abbreviation
--- a/src/HOL/Int.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Int.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -220,7 +220,7 @@
   by (transfer fixing: uminus) clarsimp
 
 lemma of_int_diff [simp]: "of_int (w - z) = of_int w - of_int z"
-by (simp add: diff_minus Groups.diff_minus)
+  using of_int_add [of w "- z"] by simp
 
 lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z"
   by (transfer fixing: times) (clarsimp simp add: algebra_simps of_nat_mult)
@@ -349,12 +349,33 @@
   shows P
   using assms by (blast dest: nat_0_le sym)
 
-lemma nat_eq_iff: "(nat w = m) = (if 0 \<le> w then w = int m else m=0)"
+lemma nat_eq_iff:
+  "nat w = m \<longleftrightarrow> (if 0 \<le> w then w = int m else m = 0)"
   by transfer (clarsimp simp add: le_imp_diff_is_add)
+ 
+corollary nat_eq_iff2:
+  "m = nat w \<longleftrightarrow> (if 0 \<le> w then w = int m else m = 0)"
+  using nat_eq_iff [of w m] by auto
+
+lemma nat_0 [simp]:
+  "nat 0 = 0"
+  by (simp add: nat_eq_iff)
 
-corollary nat_eq_iff2: "(m = nat w) = (if 0 \<le> w then w = int m else m=0)"
-by (simp only: eq_commute [of m] nat_eq_iff)
+lemma nat_1 [simp]:
+  "nat 1 = Suc 0"
+  by (simp add: nat_eq_iff)
+
+lemma nat_numeral [simp]:
+  "nat (numeral k) = numeral k"
+  by (simp add: nat_eq_iff)
 
+lemma nat_neg_numeral [simp]:
+  "nat (neg_numeral k) = 0"
+  by simp
+
+lemma nat_2: "nat 2 = Suc (Suc 0)"
+  by simp
+ 
 lemma nat_less_iff: "0 \<le> w ==> (nat w < m) = (w < of_nat m)"
   by transfer (clarsimp, arith)
 
@@ -374,12 +395,16 @@
 by (insert zless_nat_conj [of 0], auto)
 
 lemma nat_add_distrib:
-     "[| (0::int) \<le> z;  0 \<le> z' |] ==> nat (z+z') = nat z + nat z'"
+  "0 \<le> z \<Longrightarrow> 0 \<le> z' \<Longrightarrow> nat (z + z') = nat z + nat z'"
   by transfer clarsimp
 
+lemma nat_diff_distrib':
+  "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> nat (x - y) = nat x - nat y"
+  by transfer clarsimp
+ 
 lemma nat_diff_distrib:
-     "[| (0::int) \<le> z';  z' \<le> z |] ==> nat (z-z') = nat z - nat z'"
-  by transfer clarsimp
+  "0 \<le> z' \<Longrightarrow> z' \<le> z \<Longrightarrow> nat (z - z') = nat z - nat z'"
+  by (rule nat_diff_distrib') auto
 
 lemma nat_zminus_int [simp]: "nat (- int n) = 0"
   by transfer simp
@@ -399,6 +424,11 @@
 
 end
 
+lemma diff_nat_numeral [simp]: 
+  "(numeral v :: nat) - numeral v' = nat (numeral v - numeral v')"
+  by (simp only: nat_diff_distrib' zero_le_numeral nat_numeral)
+
+
 text {* For termination proofs: *}
 lemma measure_function_int[measure_function]: "is_measure (nat o abs)" ..
 
@@ -450,7 +480,7 @@
       It is proved here because attribute @{text arith_split} is not available
       in theory @{text Rings}.
       But is it really better than just rewriting with @{text abs_if}?*}
-lemma abs_split [arith_split,no_atp]:
+lemma abs_split [arith_split, no_atp]:
      "P(abs(a::'a::linordered_idom)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))"
 by (force dest: order_less_le_trans simp add: abs_if linorder_not_less)
 
@@ -722,14 +752,11 @@
 
 subsection {* Setting up simplification procedures *}
 
+lemmas of_int_simps =
+  of_int_0 of_int_1 of_int_add of_int_mult
+
 lemmas int_arith_rules =
-  neg_le_iff_le numeral_One
-  minus_zero diff_minus left_minus right_minus
-  mult_zero_left mult_zero_right mult_1_left mult_1_right
-  mult_minus_left mult_minus_right
-  minus_add_distrib minus_minus mult_assoc
-  of_nat_0 of_nat_1 of_nat_Suc of_nat_add of_nat_mult
-  of_int_0 of_int_1 of_int_add of_int_mult
+  numeral_One more_arith_simps of_nat_simps of_int_simps
 
 ML_file "Tools/int_arith.ML"
 declaration {* K Int_Arith.setup *}
@@ -768,16 +795,6 @@
 subsection{*The functions @{term nat} and @{term int}*}
 
 text{*Simplify the term @{term "w + - z"}*}
-lemmas diff_int_def_symmetric = diff_def [where 'a=int, symmetric, simp]
-
-lemma nat_0 [simp]: "nat 0 = 0"
-by (simp add: nat_eq_iff)
-
-lemma nat_1 [simp]: "nat 1 = Suc 0"
-by (subst nat_eq_iff, simp)
-
-lemma nat_2: "nat 2 = Suc (Suc 0)"
-by (subst nat_eq_iff, simp)
 
 lemma one_less_nat_eq [simp]: "(Suc 0 < nat z) = (1 < z)"
 apply (insert zless_nat_conj [of 1 z])
@@ -860,31 +877,10 @@
               if d < 0 then 0 else nat d)"
 by (simp add: Let_def nat_diff_distrib [symmetric])
 
-(* nat_diff_distrib has too-strong premises *)
-lemma nat_diff_distrib': "\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> nat (x - y) = nat x - nat y"
-apply (rule int_int_eq [THEN iffD1], clarsimp)
-apply (subst of_nat_diff)
-apply (rule nat_mono, simp_all)
-done
-
-lemma nat_numeral [simp]:
-  "nat (numeral k) = numeral k"
-  by (simp add: nat_eq_iff)
-
-lemma nat_neg_numeral [simp]:
-  "nat (neg_numeral k) = 0"
-  by simp
-
-lemma diff_nat_numeral [simp]: 
-  "(numeral v :: nat) - numeral v' = nat (numeral v - numeral v')"
-  by (simp only: nat_diff_distrib' zero_le_numeral nat_numeral)
-
 lemma nat_numeral_diff_1 [simp]:
   "numeral v - (1::nat) = nat (numeral v - 1)"
   using diff_nat_numeral [of v Num.One] by simp
 
-lemmas nat_arith = diff_nat_numeral
-
 
 subsection "Induction principles for int"
 
@@ -1158,32 +1154,32 @@
 
 text{*To Simplify Inequalities Where One Side is the Constant 1*}
 
-lemma less_minus_iff_1 [simp,no_atp]:
+lemma less_minus_iff_1 [simp]:
   fixes b::"'b::linordered_idom"
   shows "(1 < - b) = (b < -1)"
 by auto
 
-lemma le_minus_iff_1 [simp,no_atp]:
+lemma le_minus_iff_1 [simp]:
   fixes b::"'b::linordered_idom"
   shows "(1 \<le> - b) = (b \<le> -1)"
 by auto
 
-lemma equation_minus_iff_1 [simp,no_atp]:
+lemma equation_minus_iff_1 [simp]:
   fixes b::"'b::ring_1"
   shows "(1 = - b) = (b = -1)"
 by (subst equation_minus_iff, auto)
 
-lemma minus_less_iff_1 [simp,no_atp]:
+lemma minus_less_iff_1 [simp]:
   fixes a::"'b::linordered_idom"
   shows "(- a < 1) = (-1 < a)"
 by auto
 
-lemma minus_le_iff_1 [simp,no_atp]:
+lemma minus_le_iff_1 [simp]:
   fixes a::"'b::linordered_idom"
   shows "(- a \<le> 1) = (-1 \<le> a)"
 by auto
 
-lemma minus_equation_iff_1 [simp,no_atp]:
+lemma minus_equation_iff_1 [simp]:
   fixes a::"'b::ring_1"
   shows "(- a = 1) = (a = -1)"
 by (subst minus_equation_iff, auto)
@@ -1509,10 +1505,13 @@
   "sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)"
   "sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1"
   "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1"
-  unfolding sub_def dup_def numeral.simps Pos_def Neg_def
-    neg_numeral_def numeral_BitM
-  by (simp_all only: algebra_simps)
-
+  apply (simp_all only: sub_def dup_def numeral.simps Pos_def Neg_def
+    neg_numeral_def numeral_BitM)
+  apply (simp_all only: algebra_simps minus_diff_eq)
+  apply (simp_all only: add.commute [of _ "- (numeral n + numeral n)"])
+  apply (simp_all only: minus_add add.assoc left_minus)
+  apply (simp_all only: algebra_simps right_minus)
+  done
 
 text {* Implementations *}
 
--- a/src/HOL/Library/Abstract_Rat.thy	Mon Nov 11 17:34:44 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,521 +0,0 @@
-(*  Title:      HOL/Library/Abstract_Rat.thy
-    Author:     Amine Chaieb
-*)
-
-header {* Abstract rational numbers *}
-
-theory Abstract_Rat
-imports Complex_Main
-begin
-
-type_synonym Num = "int \<times> int"
-
-abbreviation Num0_syn :: Num  ("0\<^sub>N")
-  where "0\<^sub>N \<equiv> (0, 0)"
-
-abbreviation Numi_syn :: "int \<Rightarrow> Num"  ("'((_)')\<^sub>N")
-  where "(i)\<^sub>N \<equiv> (i, 1)"
-
-definition isnormNum :: "Num \<Rightarrow> bool" where
-  "isnormNum = (\<lambda>(a,b). (if a = 0 then b = 0 else b > 0 \<and> gcd a b = 1))"
-
-definition normNum :: "Num \<Rightarrow> Num" where
-  "normNum = (\<lambda>(a,b).
-    (if a=0 \<or> b = 0 then (0,0) else
-      (let g = gcd a b
-       in if b > 0 then (a div g, b div g) else (- (a div g), - (b div g)))))"
-
-declare gcd_dvd1_int[presburger] gcd_dvd2_int[presburger]
-
-lemma normNum_isnormNum [simp]: "isnormNum (normNum x)"
-proof -
-  obtain a b where x: "x = (a, b)" by (cases x)
-  { assume "a=0 \<or> b = 0" hence ?thesis by (simp add: x normNum_def isnormNum_def) }
-  moreover
-  { assume anz: "a \<noteq> 0" and bnz: "b \<noteq> 0"
-    let ?g = "gcd a b"
-    let ?a' = "a div ?g"
-    let ?b' = "b div ?g"
-    let ?g' = "gcd ?a' ?b'"
-    from anz bnz have "?g \<noteq> 0" by simp  with gcd_ge_0_int[of a b]
-    have gpos: "?g > 0" by arith
-    have gdvd: "?g dvd a" "?g dvd b" by arith+
-    from dvd_mult_div_cancel[OF gdvd(1)] dvd_mult_div_cancel[OF gdvd(2)] anz bnz
-    have nz': "?a' \<noteq> 0" "?b' \<noteq> 0" by - (rule notI, simp)+
-    from anz bnz have stupid: "a \<noteq> 0 \<or> b \<noteq> 0" by arith
-    from div_gcd_coprime_int[OF stupid] have gp1: "?g' = 1" .
-    from bnz have "b < 0 \<or> b > 0" by arith
-    moreover
-    { assume b: "b > 0"
-      from b have "?b' \<ge> 0"
-        by (presburger add: pos_imp_zdiv_nonneg_iff[OF gpos])
-      with nz' have b': "?b' > 0" by arith
-      from b b' anz bnz nz' gp1 have ?thesis
-        by (simp add: x isnormNum_def normNum_def Let_def split_def) }
-    moreover {
-      assume b: "b < 0"
-      { assume b': "?b' \<ge> 0"
-        from gpos have th: "?g \<ge> 0" by arith
-        from mult_nonneg_nonneg[OF th b'] dvd_mult_div_cancel[OF gdvd(2)]
-        have False using b by arith }
-      hence b': "?b' < 0" by (presburger add: linorder_not_le[symmetric])
-      from anz bnz nz' b b' gp1 have ?thesis
-        by (simp add: x isnormNum_def normNum_def Let_def split_def) }
-    ultimately have ?thesis by blast
-  }
-  ultimately show ?thesis by blast
-qed
-
-text {* Arithmetic over Num *}
-
-definition Nadd :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "+\<^sub>N" 60) where
-  "Nadd = (\<lambda>(a,b) (a',b'). if a = 0 \<or> b = 0 then normNum(a',b')
-    else if a'=0 \<or> b' = 0 then normNum(a,b)
-    else normNum(a*b' + b*a', b*b'))"
-
-definition Nmul :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "*\<^sub>N" 60) where
-  "Nmul = (\<lambda>(a,b) (a',b'). let g = gcd (a*a') (b*b')
-    in (a*a' div g, b*b' div g))"
-
-definition Nneg :: "Num \<Rightarrow> Num" ("~\<^sub>N")
-  where "Nneg \<equiv> (\<lambda>(a,b). (-a,b))"
-
-definition Nsub :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "-\<^sub>N" 60)
-  where "Nsub = (\<lambda>a b. a +\<^sub>N ~\<^sub>N b)"
-
-definition Ninv :: "Num \<Rightarrow> Num"
-  where "Ninv = (\<lambda>(a,b). if a < 0 then (-b, \<bar>a\<bar>) else (b,a))"
-
-definition Ndiv :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "\<div>\<^sub>N" 60)
-  where "Ndiv = (\<lambda>a b. a *\<^sub>N Ninv b)"
-
-lemma Nneg_normN[simp]: "isnormNum x \<Longrightarrow> isnormNum (~\<^sub>N x)"
-  by (simp add: isnormNum_def Nneg_def split_def)
-
-lemma Nadd_normN[simp]: "isnormNum (x +\<^sub>N y)"
-  by (simp add: Nadd_def split_def)
-
-lemma Nsub_normN[simp]: "\<lbrakk> isnormNum y\<rbrakk> \<Longrightarrow> isnormNum (x -\<^sub>N y)"
-  by (simp add: Nsub_def split_def)
-
-lemma Nmul_normN[simp]:
-  assumes xn: "isnormNum x" and yn: "isnormNum y"
-  shows "isnormNum (x *\<^sub>N y)"
-proof -
-  obtain a b where x: "x = (a, b)" by (cases x)
-  obtain a' b' where y: "y = (a', b')" by (cases y)
-  { assume "a = 0"
-    hence ?thesis using xn x y
-      by (simp add: isnormNum_def Let_def Nmul_def split_def) }
-  moreover
-  { assume "a' = 0"
-    hence ?thesis using yn x y
-      by (simp add: isnormNum_def Let_def Nmul_def split_def) }
-  moreover
-  { assume a: "a \<noteq>0" and a': "a'\<noteq>0"
-    hence bp: "b > 0" "b' > 0" using xn yn x y by (simp_all add: isnormNum_def)
-    from mult_pos_pos[OF bp] have "x *\<^sub>N y = normNum (a * a', b * b')"
-      using x y a a' bp by (simp add: Nmul_def Let_def split_def normNum_def)
-    hence ?thesis by simp }
-  ultimately show ?thesis by blast
-qed
-
-lemma Ninv_normN[simp]: "isnormNum x \<Longrightarrow> isnormNum (Ninv x)"
-  by (simp add: Ninv_def isnormNum_def split_def)
-    (cases "fst x = 0", auto simp add: gcd_commute_int)
-
-lemma isnormNum_int[simp]:
-  "isnormNum 0\<^sub>N" "isnormNum ((1::int)\<^sub>N)" "i \<noteq> 0 \<Longrightarrow> isnormNum (i)\<^sub>N"
-  by (simp_all add: isnormNum_def)
-
-
-text {* Relations over Num *}
-
-definition Nlt0:: "Num \<Rightarrow> bool"  ("0>\<^sub>N")
-  where "Nlt0 = (\<lambda>(a,b). a < 0)"
-
-definition Nle0:: "Num \<Rightarrow> bool"  ("0\<ge>\<^sub>N")
-  where "Nle0 = (\<lambda>(a,b). a \<le> 0)"
-
-definition Ngt0:: "Num \<Rightarrow> bool"  ("0<\<^sub>N")
-  where "Ngt0 = (\<lambda>(a,b). a > 0)"
-
-definition Nge0:: "Num \<Rightarrow> bool"  ("0\<le>\<^sub>N")
-  where "Nge0 = (\<lambda>(a,b). a \<ge> 0)"
-
-definition Nlt :: "Num \<Rightarrow> Num \<Rightarrow> bool"  (infix "<\<^sub>N" 55)
-  where "Nlt = (\<lambda>a b. 0>\<^sub>N (a -\<^sub>N b))"
-
-definition Nle :: "Num \<Rightarrow> Num \<Rightarrow> bool"  (infix "\<le>\<^sub>N" 55)
-  where "Nle = (\<lambda>a b. 0\<ge>\<^sub>N (a -\<^sub>N b))"
-
-definition "INum = (\<lambda>(a,b). of_int a / of_int b)"
-
-lemma INum_int [simp]: "INum (i)\<^sub>N = ((of_int i) ::'a::field)" "INum 0\<^sub>N = (0::'a::field)"
-  by (simp_all add: INum_def)
-
-lemma isnormNum_unique[simp]:
-  assumes na: "isnormNum x" and nb: "isnormNum y"
-  shows "((INum x ::'a::{field_char_0, field_inverse_zero}) = INum y) = (x = y)" (is "?lhs = ?rhs")
-proof
-  obtain a b where x: "x = (a, b)" by (cases x)
-  obtain a' b' where y: "y = (a', b')" by (cases y)
-  assume H: ?lhs
-  { assume "a = 0 \<or> b = 0 \<or> a' = 0 \<or> b' = 0"
-    hence ?rhs using na nb H
-      by (simp add: x y INum_def split_def isnormNum_def split: split_if_asm) }
-  moreover
-  { assume az: "a \<noteq> 0" and bz: "b \<noteq> 0" and a'z: "a'\<noteq>0" and b'z: "b'\<noteq>0"
-    from az bz a'z b'z na nb have pos: "b > 0" "b' > 0" by (simp_all add: x y isnormNum_def)
-    from H bz b'z have eq: "a * b' = a'*b"
-      by (simp add: x y INum_def eq_divide_eq divide_eq_eq of_int_mult[symmetric] del: of_int_mult)
-    from az a'z na nb have gcd1: "gcd a b = 1" "gcd b a = 1" "gcd a' b' = 1" "gcd b' a' = 1"
-      by (simp_all add: x y isnormNum_def add: gcd_commute_int)
-    from eq have raw_dvd: "a dvd a' * b" "b dvd b' * a" "a' dvd a * b'" "b' dvd b * a'"
-      apply -
-      apply algebra
-      apply algebra
-      apply simp
-      apply algebra
-      done
-    from zdvd_antisym_abs[OF coprime_dvd_mult_int[OF gcd1(2) raw_dvd(2)]
-        coprime_dvd_mult_int[OF gcd1(4) raw_dvd(4)]]
-      have eq1: "b = b'" using pos by arith
-      with eq have "a = a'" using pos by simp
-      with eq1 have ?rhs by (simp add: x y) }
-  ultimately show ?rhs by blast
-next
-  assume ?rhs thus ?lhs by simp
-qed
-
-
-lemma isnormNum0[simp]:
-    "isnormNum x \<Longrightarrow> (INum x = (0::'a::{field_char_0, field_inverse_zero})) = (x = 0\<^sub>N)"
-  unfolding INum_int(2)[symmetric]
-  by (rule isnormNum_unique) simp_all
-
-lemma of_int_div_aux: "d ~= 0 ==> ((of_int x)::'a::field_char_0) / (of_int d) =
-    of_int (x div d) + (of_int (x mod d)) / ((of_int d)::'a)"
-proof -
-  assume "d ~= 0"
-  let ?t = "of_int (x div d) * ((of_int d)::'a) + of_int(x mod d)"
-  let ?f = "\<lambda>x. x / of_int d"
-  have "x = (x div d) * d + x mod d"
-    by auto
-  then have eq: "of_int x = ?t"
-    by (simp only: of_int_mult[symmetric] of_int_add [symmetric])
-  then have "of_int x / of_int d = ?t / of_int d"
-    using cong[OF refl[of ?f] eq] by simp
-  then show ?thesis by (simp add: add_divide_distrib algebra_simps `d ~= 0`)
-qed
-
-lemma of_int_div: "(d::int) ~= 0 ==> d dvd n ==>
-    (of_int(n div d)::'a::field_char_0) = of_int n / of_int d"
-  apply (frule of_int_div_aux [of d n, where ?'a = 'a])
-  apply simp
-  apply (simp add: dvd_eq_mod_eq_0)
-  done
-
-
-lemma normNum[simp]: "INum (normNum x) = (INum x :: 'a::{field_char_0, field_inverse_zero})"
-proof -
-  obtain a b where x: "x = (a, b)" by (cases x)
-  { assume "a = 0 \<or> b = 0"
-    hence ?thesis by (simp add: x INum_def normNum_def split_def Let_def) }
-  moreover
-  { assume a: "a \<noteq> 0" and b: "b \<noteq> 0"
-    let ?g = "gcd a b"
-    from a b have g: "?g \<noteq> 0"by simp
-    from of_int_div[OF g, where ?'a = 'a]
-    have ?thesis by (auto simp add: x INum_def normNum_def split_def Let_def) }
-  ultimately show ?thesis by blast
-qed
-
-lemma INum_normNum_iff:
-  "(INum x ::'a::{field_char_0, field_inverse_zero}) = INum y \<longleftrightarrow> normNum x = normNum y"
-  (is "?lhs = ?rhs")
-proof -
-  have "normNum x = normNum y \<longleftrightarrow> (INum (normNum x) :: 'a) = INum (normNum y)"
-    by (simp del: normNum)
-  also have "\<dots> = ?lhs" by simp
-  finally show ?thesis by simp
-qed
-
-lemma Nadd[simp]: "INum (x +\<^sub>N y) = INum x + (INum y :: 'a :: {field_char_0, field_inverse_zero})"
-proof -
-  let ?z = "0:: 'a"
-  obtain a b where x: "x = (a, b)" by (cases x)
-  obtain a' b' where y: "y = (a', b')" by (cases y)
-  { assume "a=0 \<or> a'= 0 \<or> b =0 \<or> b' = 0"
-    hence ?thesis
-      apply (cases "a=0", simp_all add: x y Nadd_def)
-      apply (cases "b= 0", simp_all add: INum_def)
-       apply (cases "a'= 0", simp_all)
-       apply (cases "b'= 0", simp_all)
-       done }
-  moreover
-  { assume aa': "a \<noteq> 0" "a'\<noteq> 0" and bb': "b \<noteq> 0" "b' \<noteq> 0"
-    { assume z: "a * b' + b * a' = 0"
-      hence "of_int (a*b' + b*a') / (of_int b* of_int b') = ?z" by simp
-      hence "of_int b' * of_int a / (of_int b * of_int b') +
-          of_int b * of_int a' / (of_int b * of_int b') = ?z"
-        by (simp add:add_divide_distrib)
-      hence th: "of_int a / of_int b + of_int a' / of_int b' = ?z" using bb' aa'
-        by simp
-      from z aa' bb' have ?thesis
-        by (simp add: x y th Nadd_def normNum_def INum_def split_def) }
-    moreover {
-      assume z: "a * b' + b * a' \<noteq> 0"
-      let ?g = "gcd (a * b' + b * a') (b*b')"
-      have gz: "?g \<noteq> 0" using z by simp
-      have ?thesis using aa' bb' z gz
-        of_int_div[where ?'a = 'a, OF gz gcd_dvd1_int[where x="a * b' + b * a'" and y="b*b'"]]
-        of_int_div[where ?'a = 'a, OF gz gcd_dvd2_int[where x="a * b' + b * a'" and y="b*b'"]]
-        by (simp add: x y Nadd_def INum_def normNum_def Let_def add_divide_distrib) }
-    ultimately have ?thesis using aa' bb'
-      by (simp add: x y Nadd_def INum_def normNum_def Let_def) }
-  ultimately show ?thesis by blast
-qed
-
-lemma Nmul[simp]: "INum (x *\<^sub>N y) = INum x * (INum y:: 'a :: {field_char_0, field_inverse_zero})"
-proof -
-  let ?z = "0::'a"
-  obtain a b where x: "x = (a, b)" by (cases x)
-  obtain a' b' where y: "y = (a', b')" by (cases y)
-  { assume "a=0 \<or> a'= 0 \<or> b = 0 \<or> b' = 0"
-    hence ?thesis
-      apply (cases "a=0", simp_all add: x y Nmul_def INum_def Let_def)
-      apply (cases "b=0", simp_all)
-      apply (cases "a'=0", simp_all)
-      done }
-  moreover
-  { assume z: "a \<noteq> 0" "a' \<noteq> 0" "b \<noteq> 0" "b' \<noteq> 0"
-    let ?g="gcd (a*a') (b*b')"
-    have gz: "?g \<noteq> 0" using z by simp
-    from z of_int_div[where ?'a = 'a, OF gz gcd_dvd1_int[where x="a*a'" and y="b*b'"]]
-      of_int_div[where ?'a = 'a , OF gz gcd_dvd2_int[where x="a*a'" and y="b*b'"]]
-    have ?thesis by (simp add: Nmul_def x y Let_def INum_def) }
-  ultimately show ?thesis by blast
-qed
-
-lemma Nneg[simp]: "INum (~\<^sub>N x) = - (INum x ::'a:: field)"
-  by (simp add: Nneg_def split_def INum_def)
-
-lemma Nsub[simp]: "INum (x -\<^sub>N y) = INum x - (INum y:: 'a :: {field_char_0, field_inverse_zero})"
-  by (simp add: Nsub_def split_def)
-
-lemma Ninv[simp]: "INum (Ninv x) = (1::'a :: field_inverse_zero) / (INum x)"
-  by (simp add: Ninv_def INum_def split_def)
-
-lemma Ndiv[simp]: "INum (x \<div>\<^sub>N y) = INum x / (INum y ::'a :: {field_char_0, field_inverse_zero})"
-  by (simp add: Ndiv_def)
-
-lemma Nlt0_iff[simp]:
-  assumes nx: "isnormNum x"
-  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero})< 0) = 0>\<^sub>N x"
-proof -
-  obtain a b where x: "x = (a, b)" by (cases x)
-  { assume "a = 0" hence ?thesis by (simp add: x Nlt0_def INum_def) }
-  moreover
-  { assume a: "a \<noteq> 0" hence b: "(of_int b::'a) > 0"
-      using nx by (simp add: x isnormNum_def)
-    from pos_divide_less_eq[OF b, where b="of_int a" and a="0::'a"]
-    have ?thesis by (simp add: x Nlt0_def INum_def) }
-  ultimately show ?thesis by blast
-qed
-
-lemma Nle0_iff[simp]:
-  assumes nx: "isnormNum x"
-  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero}) \<le> 0) = 0\<ge>\<^sub>N x"
-proof -
-  obtain a b where x: "x = (a, b)" by (cases x)
-  { assume "a = 0" hence ?thesis by (simp add: x Nle0_def INum_def) }
-  moreover
-  { assume a: "a \<noteq> 0" hence b: "(of_int b :: 'a) > 0"
-      using nx by (simp add: x isnormNum_def)
-    from pos_divide_le_eq[OF b, where b="of_int a" and a="0::'a"]
-    have ?thesis by (simp add: x Nle0_def INum_def) }
-  ultimately show ?thesis by blast
-qed
-
-lemma Ngt0_iff[simp]:
-  assumes nx: "isnormNum x"
-  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero})> 0) = 0<\<^sub>N x"
-proof -
-  obtain a b where x: "x = (a, b)" by (cases x)
-  { assume "a = 0" hence ?thesis by (simp add: x Ngt0_def INum_def) }
-  moreover
-  { assume a: "a \<noteq> 0" hence b: "(of_int b::'a) > 0" using nx
-      by (simp add: x isnormNum_def)
-    from pos_less_divide_eq[OF b, where b="of_int a" and a="0::'a"]
-    have ?thesis by (simp add: x Ngt0_def INum_def) }
-  ultimately show ?thesis by blast
-qed
-
-lemma Nge0_iff[simp]:
-  assumes nx: "isnormNum x"
-  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero}) \<ge> 0) = 0\<le>\<^sub>N x"
-proof -
-  obtain a b where x: "x = (a, b)" by (cases x)
-  { assume "a = 0" hence ?thesis by (simp add: x Nge0_def INum_def) }
-  moreover
-  { assume "a \<noteq> 0" hence b: "(of_int b::'a) > 0" using nx
-      by (simp add: x isnormNum_def)
-    from pos_le_divide_eq[OF b, where b="of_int a" and a="0::'a"]
-    have ?thesis by (simp add: x Nge0_def INum_def) }
-  ultimately show ?thesis by blast
-qed
-
-lemma Nlt_iff[simp]:
-  assumes nx: "isnormNum x" and ny: "isnormNum y"
-  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero}) < INum y) = (x <\<^sub>N y)"
-proof -
-  let ?z = "0::'a"
-  have "((INum x ::'a) < INum y) = (INum (x -\<^sub>N y) < ?z)"
-    using nx ny by simp
-  also have "\<dots> = (0>\<^sub>N (x -\<^sub>N y))"
-    using Nlt0_iff[OF Nsub_normN[OF ny]] by simp
-  finally show ?thesis by (simp add: Nlt_def)
-qed
-
-lemma Nle_iff[simp]:
-  assumes nx: "isnormNum x" and ny: "isnormNum y"
-  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero})\<le> INum y) = (x \<le>\<^sub>N y)"
-proof -
-  have "((INum x ::'a) \<le> INum y) = (INum (x -\<^sub>N y) \<le> (0::'a))"
-    using nx ny by simp
-  also have "\<dots> = (0\<ge>\<^sub>N (x -\<^sub>N y))"
-    using Nle0_iff[OF Nsub_normN[OF ny]] by simp
-  finally show ?thesis by (simp add: Nle_def)
-qed
-
-lemma Nadd_commute:
-  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  shows "x +\<^sub>N y = y +\<^sub>N x"
-proof -
-  have n: "isnormNum (x +\<^sub>N y)" "isnormNum (y +\<^sub>N x)" by simp_all
-  have "(INum (x +\<^sub>N y)::'a) = INum (y +\<^sub>N x)" by simp
-  with isnormNum_unique[OF n] show ?thesis by simp
-qed
-
-lemma [simp]:
-  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  shows "(0, b) +\<^sub>N y = normNum y"
-    and "(a, 0) +\<^sub>N y = normNum y"
-    and "x +\<^sub>N (0, b) = normNum x"
-    and "x +\<^sub>N (a, 0) = normNum x"
-  apply (simp add: Nadd_def split_def)
-  apply (simp add: Nadd_def split_def)
-  apply (subst Nadd_commute, simp add: Nadd_def split_def)
-  apply (subst Nadd_commute, simp add: Nadd_def split_def)
-  done
-
-lemma normNum_nilpotent_aux[simp]:
-  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  assumes nx: "isnormNum x"
-  shows "normNum x = x"
-proof -
-  let ?a = "normNum x"
-  have n: "isnormNum ?a" by simp
-  have th: "INum ?a = (INum x ::'a)" by simp
-  with isnormNum_unique[OF n nx] show ?thesis by simp
-qed
-
-lemma normNum_nilpotent[simp]:
-  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  shows "normNum (normNum x) = normNum x"
-  by simp
-
-lemma normNum0[simp]: "normNum (0,b) = 0\<^sub>N" "normNum (a,0) = 0\<^sub>N"
-  by (simp_all add: normNum_def)
-
-lemma normNum_Nadd:
-  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  shows "normNum (x +\<^sub>N y) = x +\<^sub>N y" by simp
-
-lemma Nadd_normNum1[simp]:
-  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  shows "normNum x +\<^sub>N y = x +\<^sub>N y"
-proof -
-  have n: "isnormNum (normNum x +\<^sub>N y)" "isnormNum (x +\<^sub>N y)" by simp_all
-  have "INum (normNum x +\<^sub>N y) = INum x + (INum y :: 'a)" by simp
-  also have "\<dots> = INum (x +\<^sub>N y)" by simp
-  finally show ?thesis using isnormNum_unique[OF n] by simp
-qed
-
-lemma Nadd_normNum2[simp]:
-  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  shows "x +\<^sub>N normNum y = x +\<^sub>N y"
-proof -
-  have n: "isnormNum (x +\<^sub>N normNum y)" "isnormNum (x +\<^sub>N y)" by simp_all
-  have "INum (x +\<^sub>N normNum y) = INum x + (INum y :: 'a)" by simp
-  also have "\<dots> = INum (x +\<^sub>N y)" by simp
-  finally show ?thesis using isnormNum_unique[OF n] by simp
-qed
-
-lemma Nadd_assoc:
-  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  shows "x +\<^sub>N y +\<^sub>N z = x +\<^sub>N (y +\<^sub>N z)"
-proof -
-  have n: "isnormNum (x +\<^sub>N y +\<^sub>N z)" "isnormNum (x +\<^sub>N (y +\<^sub>N z))" by simp_all
-  have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" by simp
-  with isnormNum_unique[OF n] show ?thesis by simp
-qed
-
-lemma Nmul_commute: "isnormNum x \<Longrightarrow> isnormNum y \<Longrightarrow> x *\<^sub>N y = y *\<^sub>N x"
-  by (simp add: Nmul_def split_def Let_def gcd_commute_int mult_commute)
-
-lemma Nmul_assoc:
-  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  assumes nx: "isnormNum x" and ny: "isnormNum y" and nz: "isnormNum z"
-  shows "x *\<^sub>N y *\<^sub>N z = x *\<^sub>N (y *\<^sub>N z)"
-proof -
-  from nx ny nz have n: "isnormNum (x *\<^sub>N y *\<^sub>N z)" "isnormNum (x *\<^sub>N (y *\<^sub>N z))"
-    by simp_all
-  have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" by simp
-  with isnormNum_unique[OF n] show ?thesis by simp
-qed
-
-lemma Nsub0:
-  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  assumes x: "isnormNum x" and y: "isnormNum y"
-  shows "x -\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = y"
-proof -
-  fix h :: 'a
-  from isnormNum_unique[where 'a = 'a, OF Nsub_normN[OF y], where y="0\<^sub>N"]
-  have "(x -\<^sub>N y = 0\<^sub>N) = (INum (x -\<^sub>N y) = (INum 0\<^sub>N :: 'a)) " by simp
-  also have "\<dots> = (INum x = (INum y :: 'a))" by simp
-  also have "\<dots> = (x = y)" using x y by simp
-  finally show ?thesis .
-qed
-
-lemma Nmul0[simp]: "c *\<^sub>N 0\<^sub>N = 0\<^sub>N" " 0\<^sub>N *\<^sub>N c = 0\<^sub>N"
-  by (simp_all add: Nmul_def Let_def split_def)
-
-lemma Nmul_eq0[simp]:
-  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  assumes nx: "isnormNum x" and ny: "isnormNum y"
-  shows "x*\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = 0\<^sub>N \<or> y = 0\<^sub>N"
-proof -
-  fix h :: 'a
-  obtain a b where x: "x = (a, b)" by (cases x)
-  obtain a' b' where y: "y = (a', b')" by (cases y)
-  have n0: "isnormNum 0\<^sub>N" by simp
-  show ?thesis using nx ny
-    apply (simp only: isnormNum_unique[where ?'a = 'a, OF  Nmul_normN[OF nx ny] n0, symmetric]
-      Nmul[where ?'a = 'a])
-    apply (simp add: x y INum_def split_def isnormNum_def split: split_if_asm)
-    done
-qed
-
-lemma Nneg_Nneg[simp]: "~\<^sub>N (~\<^sub>N c) = c"
-  by (simp add: Nneg_def split_def)
-
-lemma Nmul1[simp]:
-    "isnormNum c \<Longrightarrow> (1)\<^sub>N *\<^sub>N c = c"
-    "isnormNum c \<Longrightarrow> c *\<^sub>N (1)\<^sub>N = c"
-  apply (simp_all add: Nmul_def Let_def split_def isnormNum_def)
-  apply (cases "fst c = 0", simp_all, cases c, simp_all)+
-  done
-
-end
--- a/src/HOL/Library/BigO.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/BigO.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -215,7 +215,7 @@
     f : lb +o O(g)"
   apply (rule set_minus_imp_plus)
   apply (rule bigo_bounded)
-  apply (auto simp add: diff_minus fun_Compl_def func_plus)
+  apply (auto simp add: fun_Compl_def func_plus)
   apply (drule_tac x = x in spec)+
   apply force
   apply (drule_tac x = x in spec)+
@@ -390,7 +390,7 @@
   apply (rule set_minus_imp_plus)
   apply (drule set_plus_imp_minus)
   apply (drule bigo_minus)
-  apply (simp add: diff_minus)
+  apply simp
   done
 
 lemma bigo_minus3: "O(-f) = O(f)"
@@ -446,7 +446,7 @@
   apply (rule bigo_minus)
   apply (subst set_minus_plus)
   apply assumption
-  apply  (simp add: diff_minus add_ac)
+  apply (simp add: add_ac)
   done
 
 lemma bigo_add_commute: "(f : g +o O(h)) = (g : f +o O(h))"
@@ -545,10 +545,9 @@
 
 lemma bigo_compose2: "f =o g +o O(h) ==> (%x. f(k x)) =o (%x. g(k x)) +o 
     O(%x. h(k x))"
-  apply (simp only: set_minus_plus [symmetric] diff_minus fun_Compl_def
-      func_plus)
-  apply (erule bigo_compose1)
-done
+  apply (simp only: set_minus_plus [symmetric] fun_Compl_def func_plus)
+  apply (drule bigo_compose1) apply (simp add: fun_diff_def)
+  done
 
 
 subsection {* Setsum *}
@@ -779,7 +778,7 @@
   apply (subst abs_of_nonneg)
   apply (drule_tac x = x in spec) back
   apply (simp add: algebra_simps)
-  apply (subst diff_minus)+
+  apply (subst diff_conv_add_uminus)+
   apply (rule add_right_mono)
   apply (erule spec)
   apply (rule order_trans) 
@@ -803,7 +802,7 @@
   apply (subst abs_of_nonneg)
   apply (drule_tac x = x in spec) back
   apply (simp add: algebra_simps)
-  apply (subst diff_minus)+
+  apply (subst diff_conv_add_uminus)+
   apply (rule add_left_mono)
   apply (rule le_imp_neg_le)
   apply (erule spec)
--- a/src/HOL/Library/Code_Target_Int.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Code_Target_Int.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -99,7 +99,7 @@
 proof -
   from mod_div_equality have *: "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp
   show ?thesis
-    by (simp add: Let_def divmod_int_mod_div mod_2_not_eq_zero_eq_one_int
+    by (simp add: Let_def divmod_int_mod_div not_mod_2_eq_0_eq_1
       of_int_add [symmetric]) (simp add: * mult_commute)
 qed
 
--- a/src/HOL/Library/ContNotDenum.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/ContNotDenum.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -131,17 +131,15 @@
 
   -- "A denotes the set of all left-most points of all the intervals ..."
   moreover obtain A where Adef: "A = ?g ` \<nat>" by simp
-  ultimately have "\<exists>x. x\<in>A"
+  ultimately have "A \<noteq> {}"
   proof -
     have "(0::nat) \<in> \<nat>" by simp
-    moreover have "?g 0 = ?g 0" by simp
-    ultimately have "?g 0 \<in> ?g ` \<nat>" by (rule  rev_image_eqI)
-    with Adef have "?g 0 \<in> A" by simp
-    thus ?thesis ..
+    with Adef show ?thesis
+      by blast
   qed
 
   -- "Now show that A is bounded above ..."
-  moreover have "\<exists>y. isUb (UNIV::real set) A y"
+  moreover have "bdd_above A"
   proof -
     {
       fix n
@@ -177,18 +175,11 @@
       obtain a and b where "f 0 = closed_int a b" and alb: "a \<le> b" by blast
     ultimately have "\<forall>n. ?g n \<in> closed_int a b" by auto
     with alb have "\<forall>n. ?g n \<le> b" using closed_int_most by blast
-    with Adef have "\<forall>y\<in>A. y\<le>b" by auto
-    hence "A *<= b" by (unfold setle_def)
-    moreover have "b \<in> (UNIV::real set)" by simp
-    ultimately have "A *<= b \<and> b \<in> (UNIV::real set)" by simp
-    hence "isUb (UNIV::real set) A b" by (unfold isUb_def)
-    thus ?thesis by auto
+    with Adef show "bdd_above A" by auto
   qed
-  -- "by the Axiom Of Completeness, A has a least upper bound ..."
-  ultimately have "\<exists>t. isLub UNIV A t" by (rule reals_complete)
 
   -- "denote this least upper bound as t ..."
-  then obtain t where tdef: "isLub UNIV A t" ..
+  def tdef: t == "Sup A"
 
   -- "and finally show that this least upper bound is in all the intervals..."
   have "\<forall>n. t \<in> f n"
@@ -229,82 +220,76 @@
         with Adef have "(?g n) \<in> A" by auto
         ultimately show ?thesis by simp
       qed 
-      with tdef show "a \<le> t" by (rule isLubD2)
+      with `bdd_above A` show "a \<le> t"
+        unfolding tdef by (intro cSup_upper)
     qed
     moreover have "t \<le> b"
-    proof -
-      have "isUb UNIV A b"
-      proof -
-        {
-          from alb int have
-            ain: "b\<in>f n \<and> (\<forall>x\<in>f n. x \<le> b)" using closed_int_most by blast
-          
-          have subsetd: "\<forall>m. \<forall>n. f (n + m) \<subseteq> f n"
-          proof (rule allI, induct_tac m)
-            show "\<forall>n. f (n + 0) \<subseteq> f n" by simp
-          next
-            fix m n
-            assume pp: "\<forall>p. f (p + n) \<subseteq> f p"
-            {
-              fix p
-              from pp have "f (p + n) \<subseteq> f p" by simp
-              moreover from subset have "f (Suc (p + n)) \<subseteq> f (p + n)" by auto
-              hence "f (p + (Suc n)) \<subseteq> f (p + n)" by simp
-              ultimately have "f (p + (Suc n)) \<subseteq> f p" by simp
-            }
-            thus "\<forall>p. f (p + Suc n) \<subseteq> f p" ..
-          qed 
-          have subsetm: "\<forall>\<alpha> \<beta>. \<alpha> \<ge> \<beta> \<longrightarrow> (f \<alpha>) \<subseteq> (f \<beta>)"
-          proof ((rule allI)+, rule impI)
-            fix \<alpha>::nat and \<beta>::nat
-            assume "\<beta> \<le> \<alpha>"
-            hence "\<exists>k. \<alpha> = \<beta> + k" by (simp only: le_iff_add)
-            then obtain k where "\<alpha> = \<beta> + k" ..
-            moreover
-            from subsetd have "f (\<beta> + k) \<subseteq> f \<beta>" by simp
-            ultimately show "f \<alpha> \<subseteq> f \<beta>" by auto
-          qed 
-          
-          fix m   
+      unfolding tdef
+    proof (rule cSup_least)
+      {
+        from alb int have
+          ain: "b\<in>f n \<and> (\<forall>x\<in>f n. x \<le> b)" using closed_int_most by blast
+        
+        have subsetd: "\<forall>m. \<forall>n. f (n + m) \<subseteq> f n"
+        proof (rule allI, induct_tac m)
+          show "\<forall>n. f (n + 0) \<subseteq> f n" by simp
+        next
+          fix m n
+          assume pp: "\<forall>p. f (p + n) \<subseteq> f p"
           {
-            assume "m \<ge> n"
-            with subsetm have "f m \<subseteq> f n" by simp
-            with ain have "\<forall>x\<in>f m. x \<le> b" by auto
-            moreover
-            from gdef have "?g m \<in> f m \<and> (\<forall>x\<in>f m. ?g m \<le> x)" by simp
-            ultimately have "?g m \<le> b" by auto
+            fix p
+            from pp have "f (p + n) \<subseteq> f p" by simp
+            moreover from subset have "f (Suc (p + n)) \<subseteq> f (p + n)" by auto
+            hence "f (p + (Suc n)) \<subseteq> f (p + n)" by simp
+            ultimately have "f (p + (Suc n)) \<subseteq> f p" by simp
           }
+          thus "\<forall>p. f (p + Suc n) \<subseteq> f p" ..
+        qed 
+        have subsetm: "\<forall>\<alpha> \<beta>. \<alpha> \<ge> \<beta> \<longrightarrow> (f \<alpha>) \<subseteq> (f \<beta>)"
+        proof ((rule allI)+, rule impI)
+          fix \<alpha>::nat and \<beta>::nat
+          assume "\<beta> \<le> \<alpha>"
+          hence "\<exists>k. \<alpha> = \<beta> + k" by (simp only: le_iff_add)
+          then obtain k where "\<alpha> = \<beta> + k" ..
+          moreover
+          from subsetd have "f (\<beta> + k) \<subseteq> f \<beta>" by simp
+          ultimately show "f \<alpha> \<subseteq> f \<beta>" by auto
+        qed 
+        
+        fix m   
+        {
+          assume "m \<ge> n"
+          with subsetm have "f m \<subseteq> f n" by simp
+          with ain have "\<forall>x\<in>f m. x \<le> b" by auto
           moreover
-          {
-            assume "\<not>(m \<ge> n)"
-            hence "m < n" by simp
-            with subsetm have sub: "(f n) \<subseteq> (f m)" by simp
-            from closed obtain ma and mb where
-              "f m = closed_int ma mb \<and> ma \<le> mb" by blast
-            hence one: "ma \<le> mb" and fm: "f m = closed_int ma mb" by auto 
-            from one alb sub fm int have "ma \<le> b" using closed_subset by blast
-            moreover have "(?g m) = ma"
-            proof -
-              from gdef have "?g m \<in> f m \<and> (\<forall>x\<in>f m. ?g m \<le> x)" ..
-              moreover from one have
-                "ma \<in> closed_int ma mb \<and> (\<forall>x\<in>closed_int ma mb. ma \<le> x)"
-                by (rule closed_int_least)
-              with fm have "ma\<in>f m \<and> (\<forall>x\<in>f m. ma \<le> x)" by simp
-              ultimately have "ma \<le> ?g m \<and> ?g m \<le> ma" by auto
-              thus "?g m = ma" by auto
-            qed
-            ultimately have "?g m \<le> b" by simp
-          } 
-          ultimately have "?g m \<le> b" by (rule case_split)
+          from gdef have "?g m \<in> f m \<and> (\<forall>x\<in>f m. ?g m \<le> x)" by simp
+          ultimately have "?g m \<le> b" by auto
         }
-        with Adef have "\<forall>y\<in>A. y\<le>b" by auto
-        hence "A *<= b" by (unfold setle_def)
-        moreover have "b \<in> (UNIV::real set)" by simp
-        ultimately have "A *<= b \<and> b \<in> (UNIV::real set)" by simp
-        thus "isUb (UNIV::real set) A b" by (unfold isUb_def)
-      qed
-      with tdef show "t \<le> b" by (rule isLub_le_isUb)
-    qed
+        moreover
+        {
+          assume "\<not>(m \<ge> n)"
+          hence "m < n" by simp
+          with subsetm have sub: "(f n) \<subseteq> (f m)" by simp
+          from closed obtain ma and mb where
+            "f m = closed_int ma mb \<and> ma \<le> mb" by blast
+          hence one: "ma \<le> mb" and fm: "f m = closed_int ma mb" by auto 
+          from one alb sub fm int have "ma \<le> b" using closed_subset by blast
+          moreover have "(?g m) = ma"
+          proof -
+            from gdef have "?g m \<in> f m \<and> (\<forall>x\<in>f m. ?g m \<le> x)" ..
+            moreover from one have
+              "ma \<in> closed_int ma mb \<and> (\<forall>x\<in>closed_int ma mb. ma \<le> x)"
+              by (rule closed_int_least)
+            with fm have "ma\<in>f m \<and> (\<forall>x\<in>f m. ma \<le> x)" by simp
+            ultimately have "ma \<le> ?g m \<and> ?g m \<le> ma" by auto
+            thus "?g m = ma" by auto
+          qed
+          ultimately have "?g m \<le> b" by simp
+        } 
+        ultimately have "?g m \<le> b" by (rule case_split)
+      }
+      with Adef show "\<And>y. y \<in> A \<Longrightarrow> y \<le> b" by auto
+    qed fact
     ultimately have "t \<in> closed_int a b" by (rule closed_mem)
     with int show "t \<in> f n" by simp
   qed
--- a/src/HOL/Library/Continuity.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Continuity.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -19,7 +19,8 @@
   "continuous F \<longleftrightarrow> (\<forall>M. chain M \<longrightarrow> F (SUP i. M i) = (SUP i. F (M i)))"
 
 lemma SUP_nat_conv:
-  "(SUP n. M n) = sup (M 0) (SUP n. M(Suc n))"
+  fixes M :: "nat \<Rightarrow> 'a::complete_lattice"
+  shows "(SUP n. M n) = sup (M 0) (SUP n. M(Suc n))"
 apply(rule order_antisym)
  apply(rule SUP_least)
  apply(case_tac n)
--- a/src/HOL/Library/Convex.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Convex.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -362,7 +362,7 @@
   shows "convex {x - y| x y. x \<in> s \<and> y \<in> t}"
 proof -
   have "{x - y| x y. x \<in> s \<and> y \<in> t} = {x + y |x y. x \<in> s \<and> y \<in> uminus ` t}"
-    unfolding diff_def by auto
+    by (auto simp add: diff_conv_add_uminus simp del: add_uminus_conv_diff)
   then show ?thesis
     using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto
 qed
--- a/src/HOL/Library/FSet.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/FSet.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -101,19 +101,25 @@
 lemma finite_Sup: "\<exists>z. finite z \<and> (\<forall>a. a \<in> X \<longrightarrow> a \<le> z) \<Longrightarrow> finite (Sup X)"
 by (auto intro: finite_subset)
 
+lemma transfer_bdd_below[transfer_rule]: "(set_rel (pcr_fset op =) ===> op =) bdd_below bdd_below"
+  by auto
+
 instance
 proof 
   fix x z :: "'a fset"
   fix X :: "'a fset set"
   {
-    assume "x \<in> X" "(\<And>a. a \<in> X \<Longrightarrow> z |\<subseteq>| a)" 
+    assume "x \<in> X" "bdd_below X" 
     then show "Inf X |\<subseteq>| x"  by transfer auto
   next
     assume "X \<noteq> {}" "(\<And>x. x \<in> X \<Longrightarrow> z |\<subseteq>| x)"
     then show "z |\<subseteq>| Inf X" by transfer (clarsimp, blast)
   next
-    assume "x \<in> X" "(\<And>a. a \<in> X \<Longrightarrow> a |\<subseteq>| z)"
-    then show "x |\<subseteq>| Sup X" by transfer (auto intro!: finite_Sup)
+    assume "x \<in> X" "bdd_above X"
+    then obtain z where "x \<in> X" "(\<And>x. x \<in> X \<Longrightarrow> x |\<subseteq>| z)"
+      by (auto simp: bdd_above_def)
+    then show "x |\<subseteq>| Sup X"
+      by transfer (auto intro!: finite_Sup)
   next
     assume "X \<noteq> {}" "(\<And>x. x \<in> X \<Longrightarrow> x |\<subseteq>| z)"
     then show "Sup X |\<subseteq>| z" by transfer (clarsimp, blast)
--- a/src/HOL/Library/Float.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Float.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -88,7 +88,7 @@
   done
 
 lemma minus_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> x - y \<in> float"
-  unfolding ab_diff_minus by (intro uminus_float plus_float)
+  using plus_float [of x "- y"] by simp
 
 lemma abs_float[simp]: "x \<in> float \<Longrightarrow> abs x \<in> float"
   by (cases x rule: linorder_cases[of 0]) auto
@@ -960,7 +960,7 @@
   also have "... < (1 / 2) * 2 powr real (rat_precision n (int x) (int y))"
     apply (rule mult_strict_right_mono) by (insert assms) auto
   also have "\<dots> = 2 powr real (rat_precision n (int x) (int y) - 1)"
-    by (simp add: powr_add diff_def powr_neg_numeral)
+    using powr_add [of 2 _ "- 1", simplified add_uminus_conv_diff] by (simp add: powr_neg_numeral)
   also have "\<dots> = 2 ^ nat (rat_precision n (int x) (int y) - 1)"
     using rat_precision_pos[of x y n] assms by (simp add: powr_realpow[symmetric])
   also have "\<dots> \<le> 2 ^ nat (rat_precision n (int x) (int y)) - 1"
--- a/src/HOL/Library/Formal_Power_Series.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Formal_Power_Series.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -231,7 +231,7 @@
 proof
   fix a b :: "'a fps"
   show "- a + a = 0" by (simp add: fps_ext)
-  show "a - b = a + - b" by (simp add: fps_ext diff_minus)
+  show "a + - b = a - b" by (simp add: fps_ext)
 qed
 
 instance fps :: (ab_group_add) ab_group_add
@@ -348,10 +348,10 @@
 instance fps :: (ring) ring ..
 
 instance fps :: (ring_1) ring_1
-  by (intro_classes, auto simp add: diff_minus distrib_right)
+  by (intro_classes, auto simp add: distrib_right)
 
 instance fps :: (comm_ring_1) comm_ring_1
-  by (intro_classes, auto simp add: diff_minus distrib_right)
+  by (intro_classes, auto simp add: distrib_right)
 
 instance fps :: (ring_no_zero_divisors) ring_no_zero_divisors
 proof
@@ -451,7 +451,7 @@
 
 definition
   dist_fps_def: "dist (a::'a fps) b =
-    (if (\<exists>n. a$n \<noteq> b$n) then inverse (2 ^ The (leastP (\<lambda>n. a$n \<noteq> b$n))) else 0)"
+    (if (\<exists>n. a$n \<noteq> b$n) then inverse (2 ^ (LEAST n. a$n \<noteq> b$n)) else 0)"
 
 lemma dist_fps_ge0: "dist (a::'a fps) b \<ge> 0"
   by (simp add: dist_fps_def)
@@ -467,34 +467,6 @@
 
 end
 
-lemma fps_nonzero_least_unique:
-  assumes a0: "a \<noteq> 0"
-  shows "\<exists>!n. leastP (\<lambda>n. a$n \<noteq> 0) n"
-proof -
-  from fps_nonzero_nth_minimal [of a] a0
-  obtain n where "a$n \<noteq> 0" "\<forall>m < n. a$m = 0" by blast
-  then have ln: "leastP (\<lambda>n. a$n \<noteq> 0) n"
-    by (auto simp add: leastP_def setge_def not_le [symmetric])
-  moreover
-  {
-    fix m
-    assume "leastP (\<lambda>n. a $ n \<noteq> 0) m"
-    then have "m = n" using ln
-      apply (auto simp add: leastP_def setge_def)
-      apply (erule allE[where x=n])
-      apply (erule allE[where x=m])
-      apply simp
-      done
-  }
-  ultimately show ?thesis by blast
-qed
-
-lemma fps_eq_least_unique:
-  assumes "(a::('a::ab_group_add) fps) \<noteq> b"
-  shows "\<exists>! n. leastP (\<lambda>n. a$n \<noteq> b$n) n"
-  using fps_nonzero_least_unique[of "a - b"] assms
-  by auto
-
 instantiation fps :: (comm_ring_1) metric_space
 begin
 
@@ -540,31 +512,15 @@
   moreover
   {
     assume ab: "a \<noteq> b" and ac: "a \<noteq> c" and bc: "b \<noteq> c"
-    let ?P = "\<lambda>a b n. a$n \<noteq> b$n"
-    from fps_eq_least_unique[OF ab] fps_eq_least_unique[OF ac]
-      fps_eq_least_unique[OF bc]
-    obtain nab nac nbc where nab: "leastP (?P a b) nab"
-      and nac: "leastP (?P a c) nac"
-      and nbc: "leastP (?P b c) nbc" by blast
-    from nab have nab': "\<And>m. m < nab \<Longrightarrow> a$m = b$m" "a$nab \<noteq> b$nab"
-      by (auto simp add: leastP_def setge_def)
-    from nac have nac': "\<And>m. m < nac \<Longrightarrow> a$m = c$m" "a$nac \<noteq> c$nac"
-      by (auto simp add: leastP_def setge_def)
-    from nbc have nbc': "\<And>m. m < nbc \<Longrightarrow> b$m = c$m" "b$nbc \<noteq> c$nbc"
-      by (auto simp add: leastP_def setge_def)
-
-    have th0: "\<And>(a::'a fps) b. a \<noteq> b \<longleftrightarrow> (\<exists>n. a$n \<noteq> b$n)"
-      by (simp add: fps_eq_iff)
-    from ab ac bc nab nac nbc
-    have dab: "dist a b = inverse (2 ^ nab)"
-      and dac: "dist a c = inverse (2 ^ nac)"
-      and dbc: "dist b c = inverse (2 ^ nbc)"
-      unfolding th0
-      apply (simp_all add: dist_fps_def)
-      apply (erule the1_equality[OF fps_eq_least_unique[OF ab]])
-      apply (erule the1_equality[OF fps_eq_least_unique[OF ac]])
-      apply (erule the1_equality[OF fps_eq_least_unique[OF bc]])
-      done
+    def n \<equiv> "\<lambda>a b::'a fps. LEAST n. a$n \<noteq> b$n"
+    then have n': "\<And>m a b. m < n a b \<Longrightarrow> a$m = b$m"
+      by (auto dest: not_less_Least)
+
+    from ab ac bc
+    have dab: "dist a b = inverse (2 ^ n a b)"
+      and dac: "dist a c = inverse (2 ^ n a c)"
+      and dbc: "dist b c = inverse (2 ^ n b c)"
+      by (simp_all add: dist_fps_def n_def fps_eq_iff)
     from ab ac bc have nz: "dist a b \<noteq> 0" "dist a c \<noteq> 0" "dist b c \<noteq> 0"
       unfolding th by simp_all
     from nz have pos: "dist a b > 0" "dist a c > 0" "dist b c > 0"
@@ -575,11 +531,13 @@
       assume h: "dist a b > dist a c + dist b c"
       then have gt: "dist a b > dist a c" "dist a b > dist b c"
         using pos by auto
-      from gt have gtn: "nab < nbc" "nab < nac"
+      from gt have gtn: "n a b < n b c" "n a b < n a c"
         unfolding dab dbc dac by (auto simp add: th1)
-      from nac'(1)[OF gtn(2)] nbc'(1)[OF gtn(1)]
-      have "a $ nab = b $ nab" by simp
-      with nab'(2) have False  by simp
+      from n'[OF gtn(2)] n'(1)[OF gtn(1)]
+      have "a $ n a b = b $ n a b" by simp
+      moreover have "a $ n a b \<noteq> b $ n a b"
+         unfolding n_def by (rule LeastI_ex) (insert ab, simp add: fps_eq_iff)
+      ultimately have False by contradiction
     }
     then have "dist a b \<le> dist a c + dist b c"
       by (auto simp add: not_le[symmetric])
@@ -649,17 +607,12 @@
       moreover
       {
         assume neq: "?s n \<noteq> a"
-        from fps_eq_least_unique[OF neq]
-        obtain k where k: "leastP (\<lambda>i. ?s n $ i \<noteq> a$i) k" by blast
-        have th0: "\<And>(a::'a fps) b. a \<noteq> b \<longleftrightarrow> (\<exists>n. a$n \<noteq> b$n)"
-          by (simp add: fps_eq_iff)
+        def k \<equiv> "LEAST i. ?s n $ i \<noteq> a $ i"
         from neq have dth: "dist (?s n) a = (1/2)^k"
-          unfolding th0 dist_fps_def
-          unfolding the1_equality[OF fps_eq_least_unique[OF neq], OF k]
-          by (auto simp add: inverse_eq_divide power_divide)
-
-        from k have kn: "k > n"
-          by (simp add: leastP_def setge_def fps_sum_rep_nth split:split_if_asm)
+          by (auto simp add: dist_fps_def inverse_eq_divide power_divide k_def fps_eq_iff)
+
+        from neq have kn: "k > n"
+          by (auto simp: fps_sum_rep_nth not_le k_def fps_eq_iff split: split_if_asm intro: LeastI2_ex)
         then have "dist (?s n) a < (1/2)^n" unfolding dth
           by (auto intro: power_strict_decreasing)
         also have "\<dots> <= (1/2)^n0" using nn0
@@ -888,7 +841,7 @@
   using fps_deriv_linear[of 1 f 1 g] by simp
 
 lemma fps_deriv_sub[simp]: "fps_deriv ((f:: ('a::comm_ring_1) fps) - g) = fps_deriv f - fps_deriv g"
-  unfolding diff_minus by simp
+  using fps_deriv_add [of f "- g"] by simp
 
 lemma fps_deriv_const[simp]: "fps_deriv (fps_const c) = 0"
   by (simp add: fps_ext fps_deriv_def fps_const_def)
@@ -978,7 +931,7 @@
 
 lemma fps_nth_deriv_sub[simp]:
   "fps_nth_deriv n ((f:: ('a::comm_ring_1) fps) - g) = fps_nth_deriv n f - fps_nth_deriv n g"
-  unfolding diff_minus fps_nth_deriv_add by simp
+  using fps_nth_deriv_add [of n f "- g"] by simp
 
 lemma fps_nth_deriv_0[simp]: "fps_nth_deriv n 0 = 0"
   by (induct n) simp_all
@@ -2634,7 +2587,7 @@
   by (simp add: fps_eq_iff fps_compose_nth field_simps setsum_negf[symmetric])
 
 lemma fps_compose_sub_distrib: "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)"
-  unfolding diff_minus fps_compose_uminus fps_compose_add_distrib ..
+  using fps_compose_add_distrib [of a "- b" c] by (simp add: fps_compose_uminus)
 
 lemma X_fps_compose: "X oo a = Abs_fps (\<lambda>n. if n = 0 then (0::'a::comm_ring_1) else a$n)"
   by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum_delta)
@@ -3797,20 +3750,14 @@
   assumes "dist f g < inverse (2 ^ i)"
     and"j \<le> i"
   shows "f $ j = g $ j"
-proof (cases "f = g")
-  case False
-  hence "\<exists>n. f $ n \<noteq> g $ n" by (simp add: fps_eq_iff)
-  with assms have "i < The (leastP (\<lambda>n. f $ n \<noteq> g $ n))"
+proof (rule ccontr)
+  assume "f $ j \<noteq> g $ j"
+  then have "\<exists>n. f $ n \<noteq> g $ n" by auto
+  with assms have "i < (LEAST n. f $ n \<noteq> g $ n)"
     by (simp add: split_if_asm dist_fps_def)
-  moreover
-  from fps_eq_least_unique[OF `f \<noteq> g`]
-  obtain n where n: "leastP (\<lambda>n. f$n \<noteq> g$n) n" "The (leastP (\<lambda>n. f $ n \<noteq> g $ n)) = n" by blast
-  moreover from n have "\<And>m. m < n \<Longrightarrow> f$m = g$m" "f$n \<noteq> g$n"
-    by (auto simp add: leastP_def setge_def)
-  ultimately show ?thesis using `j \<le> i` by simp
-next
-  case True
-  then show ?thesis by simp
+  also have "\<dots> \<le> j"
+    using `f $ j \<noteq> g $ j` by (auto intro: Least_le)
+  finally show False using `j \<le> i` by simp
 qed
 
 lemma nth_equal_imp_dist_less:
@@ -3819,18 +3766,13 @@
 proof (cases "f = g")
   case False
   hence "\<exists>n. f $ n \<noteq> g $ n" by (simp add: fps_eq_iff)
-  with assms have "dist f g = inverse (2 ^ (The (leastP (\<lambda>n. f $ n \<noteq> g $ n))))"
+  with assms have "dist f g = inverse (2 ^ (LEAST n. f $ n \<noteq> g $ n))"
     by (simp add: split_if_asm dist_fps_def)
   moreover
-  from fps_eq_least_unique[OF `f \<noteq> g`]
-  obtain n where "leastP (\<lambda>n. f$n \<noteq> g$n) n" "The (leastP (\<lambda>n. f $ n \<noteq> g $ n)) = n" by blast
-  with assms have "i < The (leastP (\<lambda>n. f $ n \<noteq> g $ n))"
-    by (metis (full_types) leastPD1 not_le)
+  from assms `\<exists>n. f $ n \<noteq> g $ n` have "i < (LEAST n. f $ n \<noteq> g $ n)"
+    by (metis (mono_tags) LeastI not_less)
   ultimately show ?thesis by simp
-next
-  case True
-  then show ?thesis by simp
-qed
+qed simp
 
 lemma dist_less_eq_nth_equal: "dist f g < inverse (2 ^ i) \<longleftrightarrow> (\<forall>j \<le> i. f $ j = g $ j)"
   using dist_less_imp_nth_equal nth_equal_imp_dist_less by blast
--- a/src/HOL/Library/Fraction_Field.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Fraction_Field.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -132,7 +132,7 @@
 lemma diff_fract [simp]:
   assumes "b \<noteq> 0" and "d \<noteq> 0"
   shows "Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
-  using assms by (simp add: diff_fract_def diff_minus)
+  using assms by (simp add: diff_fract_def)
 
 definition mult_fract_def:
   "q * r = Abs_fract (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
--- a/src/HOL/Library/Function_Algebras.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Function_Algebras.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -83,10 +83,10 @@
 
 instance "fun" :: (type, group_add) group_add
   by default
-    (simp_all add: fun_eq_iff diff_minus)
+    (simp_all add: fun_eq_iff)
 
 instance "fun" :: (type, ab_group_add) ab_group_add
-  by default (simp_all add: diff_minus)
+  by default simp_all
 
 
 text {* Multiplicative structures *}
--- a/src/HOL/Library/Fundamental_Theorem_Algebra.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Fundamental_Theorem_Algebra.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -156,28 +156,11 @@
 text{* An alternative useful formulation of completeness of the reals *}
 lemma real_sup_exists: assumes ex: "\<exists>x. P x" and bz: "\<exists>z. \<forall>x. P x \<longrightarrow> x < z"
   shows "\<exists>(s::real). \<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < s"
-proof-
-  from ex bz obtain x Y where x: "P x" and Y: "\<And>x. P x \<Longrightarrow> x < Y"  by blast
-  from ex have thx:"\<exists>x. x \<in> Collect P" by blast
-  from bz have thY: "\<exists>Y. isUb UNIV (Collect P) Y"
-    by(auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def order_le_less)
-  from reals_complete[OF thx thY] obtain L where L: "isLub UNIV (Collect P) L"
-    by blast
-  from Y[OF x] have xY: "x < Y" .
-  from L have L': "\<forall>x. P x \<longrightarrow> x \<le> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
-  from Y have Y': "\<forall>x. P x \<longrightarrow> x \<le> Y"
-    apply (clarsimp, atomize (full)) by auto
-  from L Y' have "L \<le> Y" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
-  {fix y
-    {fix z assume z: "P z" "y < z"
-      from L' z have "y < L" by auto }
-    moreover
-    {assume yL: "y < L" "\<forall>z. P z \<longrightarrow> \<not> y < z"
-      hence nox: "\<forall>z. P z \<longrightarrow> y \<ge> z" by auto
-      from nox L have "y \<ge> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
-      with yL(1) have False  by arith}
-    ultimately have "(\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < L" by blast}
-  thus ?thesis by blast
+proof
+  from bz have "bdd_above (Collect P)"
+    by (force intro: less_imp_le)
+  then show "\<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < Sup (Collect P)"
+    using ex bz by (subst less_cSup_iff) auto
 qed
 
 subsection {* Fundamental theorem of algebra *}
@@ -224,12 +207,12 @@
     from unimodular_reduce_norm[OF th0] o
     have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
       apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1", rule_tac x="1" in exI, simp)
-      apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp add: diff_minus)
+      apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp del: minus_one add: minus_one [symmetric])
       apply (cases "cmod (complex_of_real (cmod b) / b + ii) < 1")
       apply (cases "even m", rule_tac x="ii" in exI, simp add: m power_mult)
       apply (rule_tac x="- ii" in exI, simp add: m power_mult)
-      apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult diff_minus)
-      apply (rule_tac x="ii" in exI, simp add: m power_mult diff_minus)
+      apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult)
+      apply (rule_tac x="ii" in exI, simp add: m power_mult)
       done
     then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1" by blast
     let ?w = "v / complex_of_real (root n (cmod b))"
@@ -954,7 +937,7 @@
 
 lemma mpoly_sub_conv:
   "poly p (x::complex) - poly q x \<equiv> poly p x + -1 * poly q x"
-  by (simp add: diff_minus)
+  by simp
 
 lemma poly_pad_rule: "poly p x = 0 ==> poly (pCons 0 p) x = (0::complex)" by simp
 
--- a/src/HOL/Library/Glbs.thy	Mon Nov 11 17:34:44 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,79 +0,0 @@
-(* Author: Amine Chaieb, University of Cambridge *)
-
-header {* Definitions of Lower Bounds and Greatest Lower Bounds, analogous to Lubs *}
-
-theory Glbs
-imports Lubs
-begin
-
-definition greatestP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a::ord \<Rightarrow> bool"
-  where "greatestP P x = (P x \<and> Collect P *<=  x)"
-
-definition isLb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
-  where "isLb R S x = (x <=* S \<and> x: R)"
-
-definition isGlb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
-  where "isGlb R S x = greatestP (isLb R S) x"
-
-definition lbs :: "'a set \<Rightarrow> 'a::ord set \<Rightarrow> 'a set"
-  where "lbs R S = Collect (isLb R S)"
-
-
-subsection {* Rules about the Operators @{term greatestP}, @{term isLb}
-  and @{term isGlb} *}
-
-lemma leastPD1: "greatestP P x \<Longrightarrow> P x"
-  by (simp add: greatestP_def)
-
-lemma greatestPD2: "greatestP P x \<Longrightarrow> Collect P *<= x"
-  by (simp add: greatestP_def)
-
-lemma greatestPD3: "greatestP P x \<Longrightarrow> y: Collect P \<Longrightarrow> x \<ge> y"
-  by (blast dest!: greatestPD2 setleD)
-
-lemma isGlbD1: "isGlb R S x \<Longrightarrow> x <=* S"
-  by (simp add: isGlb_def isLb_def greatestP_def)
-
-lemma isGlbD1a: "isGlb R S x \<Longrightarrow> x: R"
-  by (simp add: isGlb_def isLb_def greatestP_def)
-
-lemma isGlb_isLb: "isGlb R S x \<Longrightarrow> isLb R S x"
-  unfolding isLb_def by (blast dest: isGlbD1 isGlbD1a)
-
-lemma isGlbD2: "isGlb R S x \<Longrightarrow> y : S \<Longrightarrow> y \<ge> x"
-  by (blast dest!: isGlbD1 setgeD)
-
-lemma isGlbD3: "isGlb R S x \<Longrightarrow> greatestP (isLb R S) x"
-  by (simp add: isGlb_def)
-
-lemma isGlbI1: "greatestP (isLb R S) x \<Longrightarrow> isGlb R S x"
-  by (simp add: isGlb_def)
-
-lemma isGlbI2: "isLb R S x \<Longrightarrow> Collect (isLb R S) *<= x \<Longrightarrow> isGlb R S x"
-  by (simp add: isGlb_def greatestP_def)
-
-lemma isLbD: "isLb R S x \<Longrightarrow> y : S \<Longrightarrow> y \<ge> x"
-  by (simp add: isLb_def setge_def)
-
-lemma isLbD2: "isLb R S x \<Longrightarrow> x <=* S "
-  by (simp add: isLb_def)
-
-lemma isLbD2a: "isLb R S x \<Longrightarrow> x: R"
-  by (simp add: isLb_def)
-
-lemma isLbI: "x <=* S \<Longrightarrow> x: R \<Longrightarrow> isLb R S x"
-  by (simp add: isLb_def)
-
-lemma isGlb_le_isLb: "isGlb R S x \<Longrightarrow> isLb R S y \<Longrightarrow> x \<ge> y"
-  unfolding isGlb_def by (blast intro!: greatestPD3)
-
-lemma isGlb_ubs: "isGlb R S x \<Longrightarrow> lbs R S *<= x"
-  unfolding lbs_def isGlb_def by (rule greatestPD2)
-
-lemma isGlb_unique: "[| isGlb R S x; isGlb R S y |] ==> x = (y::'a::linorder)"
-  apply (frule isGlb_isLb)
-  apply (frule_tac x = y in isGlb_isLb)
-  apply (blast intro!: order_antisym dest!: isGlb_le_isLb)
-  done
-
-end
--- a/src/HOL/Library/Inner_Product.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Inner_Product.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -41,7 +41,7 @@
   using inner_add_left [of x "- x" y] by simp
 
 lemma inner_diff_left: "inner (x - y) z = inner x z - inner y z"
-  by (simp add: diff_minus inner_add_left)
+  using inner_add_left [of x "- y" z] by simp
 
 lemma inner_setsum_left: "inner (\<Sum>x\<in>A. f x) y = (\<Sum>x\<in>A. inner (f x) y)"
   by (cases "finite A", induct set: finite, simp_all add: inner_add_left)
--- a/src/HOL/Library/Lattice_Algebras.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Lattice_Algebras.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -13,9 +13,7 @@
   apply (rule antisym)
   apply (simp_all add: le_infI)
   apply (rule add_le_imp_le_left [of "uminus a"])
-  apply (simp only: add_assoc [symmetric], simp)
-  apply rule
-  apply (rule add_le_imp_le_left[of "a"], simp only: add_assoc[symmetric], simp)+
+  apply (simp only: add_assoc [symmetric], simp add: diff_le_eq add.commute)
   done
 
 lemma add_inf_distrib_right: "inf a b + c = inf (a + c) (b + c)"
@@ -33,11 +31,10 @@
 lemma add_sup_distrib_left: "a + sup b c = sup (a + b) (a + c)"
   apply (rule antisym)
   apply (rule add_le_imp_le_left [of "uminus a"])
-  apply (simp only: add_assoc[symmetric], simp)
-  apply rule
+  apply (simp only: add_assoc [symmetric], simp)
+  apply (simp add: le_diff_eq add.commute)
+  apply (rule le_supI)
   apply (rule add_le_imp_le_left [of "a"], simp only: add_assoc[symmetric], simp)+
-  apply (rule le_supI)
-  apply (simp_all)
   done
 
 lemma add_sup_distrib_right: "sup a b + c = sup (a+c) (b+c)"
@@ -87,9 +84,15 @@
 lemma neg_inf_eq_sup: "- inf a b = sup (-a) (-b)"
   by (simp add: inf_eq_neg_sup)
 
+lemma diff_inf_eq_sup: "a - inf b c = a + sup (- b) (- c)"
+  using neg_inf_eq_sup [of b c, symmetric] by simp
+
 lemma neg_sup_eq_inf: "- sup a b = inf (-a) (-b)"
   by (simp add: sup_eq_neg_inf)
 
+lemma diff_sup_eq_inf: "a - sup b c = a + inf (- b) (- c)"
+  using neg_sup_eq_inf [of b c, symmetric] by simp
+
 lemma add_eq_inf_sup: "a + b = sup a b + inf a b"
 proof -
   have "0 = - inf 0 (a-b) + inf (a-b) 0"
@@ -97,8 +100,8 @@
   hence "0 = sup 0 (b-a) + inf (a-b) 0"
     by (simp add: inf_eq_neg_sup)
   hence "0 = (-a + sup a b) + (inf a b + (-b))"
-    by (simp add: add_sup_distrib_left add_inf_distrib_right) (simp add: algebra_simps)
-  thus ?thesis by (simp add: algebra_simps)
+    by (simp only: add_sup_distrib_left add_inf_distrib_right) simp
+  then show ?thesis by (simp add: algebra_simps)
 qed
 
 
@@ -251,7 +254,7 @@
     apply assumption
     apply (rule notI)
     unfolding double_zero [symmetric, of a]
-    apply simp
+    apply blast
     done
 qed
 
@@ -259,7 +262,8 @@
   "a + a \<le> 0 \<longleftrightarrow> a \<le> 0"
 proof -
   have "a + a \<le> 0 \<longleftrightarrow> 0 \<le> - (a + a)" by (subst le_minus_iff, simp)
-  moreover have "\<dots> \<longleftrightarrow> a \<le> 0" by simp
+  moreover have "\<dots> \<longleftrightarrow> a \<le> 0"
+    by (simp only: minus_add_distrib zero_le_double_add_iff_zero_le_single_add) simp
   ultimately show ?thesis by blast
 qed
 
@@ -267,11 +271,12 @@
   "a + a < 0 \<longleftrightarrow> a < 0"
 proof -
   have "a + a < 0 \<longleftrightarrow> 0 < - (a + a)" by (subst less_minus_iff, simp)
-  moreover have "\<dots> \<longleftrightarrow> a < 0" by simp
+  moreover have "\<dots> \<longleftrightarrow> a < 0"
+    by (simp only: minus_add_distrib zero_less_double_add_iff_zero_less_single_add) simp
   ultimately show ?thesis by blast
 qed
 
-declare neg_inf_eq_sup [simp] neg_sup_eq_inf [simp]
+declare neg_inf_eq_sup [simp] neg_sup_eq_inf [simp] diff_inf_eq_sup [simp] diff_sup_eq_inf [simp]
 
 lemma le_minus_self_iff: "a \<le> - a \<longleftrightarrow> a \<le> 0"
 proof -
@@ -326,7 +331,7 @@
   then have "0 \<le> sup a (- a)" unfolding abs_lattice .
   then have "sup (sup a (- a)) 0 = sup a (- a)" by (rule sup_absorb1)
   then show ?thesis
-    by (simp add: add_sup_inf_distribs sup_aci pprt_def nprt_def diff_minus abs_lattice)
+    by (simp add: add_sup_inf_distribs ac_simps pprt_def nprt_def abs_lattice)
 qed
 
 subclass ordered_ab_group_add_abs
@@ -355,16 +360,17 @@
   show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
   proof -
     have g:"abs a + abs b = sup (a+b) (sup (-a-b) (sup (-a+b) (a + (-b))))" (is "_=sup ?m ?n")
-      by (simp add: abs_lattice add_sup_inf_distribs sup_aci diff_minus)
+      by (simp add: abs_lattice add_sup_inf_distribs sup_aci ac_simps)
     have a: "a + b <= sup ?m ?n" by simp
     have b: "- a - b <= ?n" by simp
     have c: "?n <= sup ?m ?n" by simp
     from b c have d: "-a-b <= sup ?m ?n" by (rule order_trans)
-    have e:"-a-b = -(a+b)" by (simp add: diff_minus)
+    have e:"-a-b = -(a+b)" by simp
     from a d e have "abs(a+b) <= sup ?m ?n"
       apply -
       apply (drule abs_leI)
-      apply auto
+      apply (simp_all only: algebra_simps ac_simps minus_add)
+      apply (metis add_uminus_conv_diff d sup_commute uminus_add_conv_diff)
       done
     with g[symmetric] show ?thesis by simp
   qed
@@ -421,14 +427,12 @@
   }
   note b = this[OF refl[of a] refl[of b]]
   have xy: "- ?x <= ?y"
-    apply (simp)
-    apply (rule order_trans [OF add_nonpos_nonpos add_nonneg_nonneg])
-    apply (simp_all add: mult_nonneg_nonneg mult_nonpos_nonpos)
+    apply simp
+    apply (metis (full_types) add_increasing add_uminus_conv_diff lattice_ab_group_add_class.minus_le_self_iff minus_add_distrib mult_nonneg_nonneg mult_nonpos_nonpos nprt_le_zero zero_le_pprt)
     done
   have yx: "?y <= ?x"
-    apply (simp add:diff_minus)
-    apply (rule order_trans [OF add_nonpos_nonpos add_nonneg_nonneg])
-    apply (simp_all add: mult_nonneg_nonpos mult_nonpos_nonneg)
+    apply simp
+    apply (metis (full_types) add_nonpos_nonpos add_uminus_conv_diff lattice_ab_group_add_class.le_minus_self_iff minus_add_distrib mult_nonneg_nonpos mult_nonpos_nonneg nprt_le_zero zero_le_pprt)
     done
   have i1: "a*b <= abs a * abs b" by (simp only: a b yx)
   have i2: "- (abs a * abs b) <= a*b" by (simp only: a b xy)
@@ -549,7 +553,7 @@
     by simp
   then have "-(- nprt a1 * pprt b2 + - nprt a2 * nprt b2 + - pprt a1 * pprt b1 + - pprt a2 * nprt b1) <= a * b"
     by (simp only: minus_le_iff)
-  then show ?thesis by simp
+  then show ?thesis by (simp add: algebra_simps)
 qed
 
 instance int :: lattice_ring
@@ -567,3 +571,4 @@
 qed
 
 end
+
--- a/src/HOL/Library/Library.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Library.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1,7 +1,6 @@
 (*<*)
 theory Library
 imports
-  Abstract_Rat
   AList
   BigO
   Binomial
@@ -65,7 +64,6 @@
   Sublist
   Sum_of_Squares
   Transitive_Closure_Table
-  Univ_Poly
   Wfrec
   While_Combinator
   Zorn
--- a/src/HOL/Library/Liminf_Limsup.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Liminf_Limsup.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -21,19 +21,21 @@
   by (blast intro: less_imp_le less_trans le_less_trans dest: dense)
 
 lemma SUPR_pair:
-  "(SUP i : A. SUP j : B. f i j) = (SUP p : A \<times> B. f (fst p) (snd p))"
+  fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: complete_lattice"
+  shows "(SUP i : A. SUP j : B. f i j) = (SUP p : A \<times> B. f (fst p) (snd p))"
   by (rule antisym) (auto intro!: SUP_least SUP_upper2)
 
 lemma INFI_pair:
-  "(INF i : A. INF j : B. f i j) = (INF p : A \<times> B. f (fst p) (snd p))"
+  fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: complete_lattice"
+  shows "(INF i : A. INF j : B. f i j) = (INF p : A \<times> B. f (fst p) (snd p))"
   by (rule antisym) (auto intro!: INF_greatest INF_lower2)
 
 subsubsection {* @{text Liminf} and @{text Limsup} *}
 
-definition
+definition Liminf :: "'a filter \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b :: complete_lattice" where
   "Liminf F f = (SUP P:{P. eventually P F}. INF x:{x. P x}. f x)"
 
-definition
+definition Limsup :: "'a filter \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b :: complete_lattice" where
   "Limsup F f = (INF P:{P. eventually P F}. SUP x:{x. P x}. f x)"
 
 abbreviation "liminf \<equiv> Liminf sequentially"
@@ -50,21 +52,17 @@
     (\<And>y. (\<And>P. eventually P F \<Longrightarrow> y \<le> SUPR (Collect P) f) \<Longrightarrow> y \<le> x) \<Longrightarrow> Limsup F f = x"
   unfolding Limsup_def by (auto intro!: INF_eqI)
 
-lemma liminf_SUPR_INFI:
-  fixes f :: "nat \<Rightarrow> 'a :: complete_lattice"
-  shows "liminf f = (SUP n. INF m:{n..}. f m)"
+lemma liminf_SUPR_INFI: "liminf f = (SUP n. INF m:{n..}. f m)"
   unfolding Liminf_def eventually_sequentially
   by (rule SUPR_eq) (auto simp: atLeast_def intro!: INF_mono)
 
-lemma limsup_INFI_SUPR:
-  fixes f :: "nat \<Rightarrow> 'a :: complete_lattice"
-  shows "limsup f = (INF n. SUP m:{n..}. f m)"
+lemma limsup_INFI_SUPR: "limsup f = (INF n. SUP m:{n..}. f m)"
   unfolding Limsup_def eventually_sequentially
   by (rule INFI_eq) (auto simp: atLeast_def intro!: SUP_mono)
 
 lemma Limsup_const: 
   assumes ntriv: "\<not> trivial_limit F"
-  shows "Limsup F (\<lambda>x. c) = (c::'a::complete_lattice)"
+  shows "Limsup F (\<lambda>x. c) = c"
 proof -
   have *: "\<And>P. Ex P \<longleftrightarrow> P \<noteq> (\<lambda>x. False)" by auto
   have "\<And>P. eventually P F \<Longrightarrow> (SUP x : {x. P x}. c) = c"
@@ -77,7 +75,7 @@
 
 lemma Liminf_const:
   assumes ntriv: "\<not> trivial_limit F"
-  shows "Liminf F (\<lambda>x. c) = (c::'a::complete_lattice)"
+  shows "Liminf F (\<lambda>x. c) = c"
 proof -
   have *: "\<And>P. Ex P \<longleftrightarrow> P \<noteq> (\<lambda>x. False)" by auto
   have "\<And>P. eventually P F \<Longrightarrow> (INF x : {x. P x}. c) = c"
@@ -89,7 +87,6 @@
 qed
 
 lemma Liminf_mono:
-  fixes f g :: "'a => 'b :: complete_lattice"
   assumes ev: "eventually (\<lambda>x. f x \<le> g x) F"
   shows "Liminf F f \<le> Liminf F g"
   unfolding Liminf_def
@@ -101,13 +98,11 @@
 qed
 
 lemma Liminf_eq:
-  fixes f g :: "'a \<Rightarrow> 'b :: complete_lattice"
   assumes "eventually (\<lambda>x. f x = g x) F"
   shows "Liminf F f = Liminf F g"
   by (intro antisym Liminf_mono eventually_mono[OF _ assms]) auto
 
 lemma Limsup_mono:
-  fixes f g :: "'a \<Rightarrow> 'b :: complete_lattice"
   assumes ev: "eventually (\<lambda>x. f x \<le> g x) F"
   shows "Limsup F f \<le> Limsup F g"
   unfolding Limsup_def
@@ -119,18 +114,16 @@
 qed
 
 lemma Limsup_eq:
-  fixes f g :: "'a \<Rightarrow> 'b :: complete_lattice"
   assumes "eventually (\<lambda>x. f x = g x) net"
   shows "Limsup net f = Limsup net g"
   by (intro antisym Limsup_mono eventually_mono[OF _ assms]) auto
 
 lemma Liminf_le_Limsup:
-  fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
   assumes ntriv: "\<not> trivial_limit F"
   shows "Liminf F f \<le> Limsup F f"
   unfolding Limsup_def Liminf_def
-  apply (rule complete_lattice_class.SUP_least)
-  apply (rule complete_lattice_class.INF_greatest)
+  apply (rule SUP_least)
+  apply (rule INF_greatest)
 proof safe
   fix P Q assume "eventually P F" "eventually Q F"
   then have "eventually (\<lambda>x. P x \<and> Q x) F" (is "eventually ?C F") by (rule eventually_conj)
@@ -146,14 +139,12 @@
 qed
 
 lemma Liminf_bounded:
-  fixes X Y :: "'a \<Rightarrow> 'b::complete_lattice"
   assumes ntriv: "\<not> trivial_limit F"
   assumes le: "eventually (\<lambda>n. C \<le> X n) F"
   shows "C \<le> Liminf F X"
   using Liminf_mono[OF le] Liminf_const[OF ntriv, of C] by simp
 
 lemma Limsup_bounded:
-  fixes X Y :: "'a \<Rightarrow> 'b::complete_lattice"
   assumes ntriv: "\<not> trivial_limit F"
   assumes le: "eventually (\<lambda>n. X n \<le> C) F"
   shows "Limsup F X \<le> C"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Lubs_Glbs.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -0,0 +1,245 @@
+(*  Title:      HOL/Library/Lubs_Glbs.thy
+    Author:     Jacques D. Fleuriot, University of Cambridge
+    Author:     Amine Chaieb, University of Cambridge *)
+
+header {* Definitions of Least Upper Bounds and Greatest Lower Bounds *}
+
+theory Lubs_Glbs
+imports Complex_Main
+begin
+
+text {* Thanks to suggestions by James Margetson *}
+
+definition setle :: "'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"  (infixl "*<=" 70)
+  where "S *<= x = (ALL y: S. y \<le> x)"
+
+definition setge :: "'a::ord \<Rightarrow> 'a set \<Rightarrow> bool"  (infixl "<=*" 70)
+  where "x <=* S = (ALL y: S. x \<le> y)"
+
+
+subsection {* Rules for the Relations @{text "*<="} and @{text "<=*"} *}
+
+lemma setleI: "ALL y: S. y \<le> x \<Longrightarrow> S *<= x"
+  by (simp add: setle_def)
+
+lemma setleD: "S *<= x \<Longrightarrow> y: S \<Longrightarrow> y \<le> x"
+  by (simp add: setle_def)
+
+lemma setgeI: "ALL y: S. x \<le> y \<Longrightarrow> x <=* S"
+  by (simp add: setge_def)
+
+lemma setgeD: "x <=* S \<Longrightarrow> y: S \<Longrightarrow> x \<le> y"
+  by (simp add: setge_def)
+
+
+definition leastP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a::ord \<Rightarrow> bool"
+  where "leastP P x = (P x \<and> x <=* Collect P)"
+
+definition isUb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
+  where "isUb R S x = (S *<= x \<and> x: R)"
+
+definition isLub :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
+  where "isLub R S x = leastP (isUb R S) x"
+
+definition ubs :: "'a set \<Rightarrow> 'a::ord set \<Rightarrow> 'a set"
+  where "ubs R S = Collect (isUb R S)"
+
+
+subsection {* Rules about the Operators @{term leastP}, @{term ub} and @{term lub} *}
+
+lemma leastPD1: "leastP P x \<Longrightarrow> P x"
+  by (simp add: leastP_def)
+
+lemma leastPD2: "leastP P x \<Longrightarrow> x <=* Collect P"
+  by (simp add: leastP_def)
+
+lemma leastPD3: "leastP P x \<Longrightarrow> y: Collect P \<Longrightarrow> x \<le> y"
+  by (blast dest!: leastPD2 setgeD)
+
+lemma isLubD1: "isLub R S x \<Longrightarrow> S *<= x"
+  by (simp add: isLub_def isUb_def leastP_def)
+
+lemma isLubD1a: "isLub R S x \<Longrightarrow> x: R"
+  by (simp add: isLub_def isUb_def leastP_def)
+
+lemma isLub_isUb: "isLub R S x \<Longrightarrow> isUb R S x"
+  unfolding isUb_def by (blast dest: isLubD1 isLubD1a)
+
+lemma isLubD2: "isLub R S x \<Longrightarrow> y : S \<Longrightarrow> y \<le> x"
+  by (blast dest!: isLubD1 setleD)
+
+lemma isLubD3: "isLub R S x \<Longrightarrow> leastP (isUb R S) x"
+  by (simp add: isLub_def)
+
+lemma isLubI1: "leastP(isUb R S) x \<Longrightarrow> isLub R S x"
+  by (simp add: isLub_def)
+
+lemma isLubI2: "isUb R S x \<Longrightarrow> x <=* Collect (isUb R S) \<Longrightarrow> isLub R S x"
+  by (simp add: isLub_def leastP_def)
+
+lemma isUbD: "isUb R S x \<Longrightarrow> y : S \<Longrightarrow> y \<le> x"
+  by (simp add: isUb_def setle_def)
+
+lemma isUbD2: "isUb R S x \<Longrightarrow> S *<= x"
+  by (simp add: isUb_def)
+
+lemma isUbD2a: "isUb R S x \<Longrightarrow> x: R"
+  by (simp add: isUb_def)
+
+lemma isUbI: "S *<= x \<Longrightarrow> x: R \<Longrightarrow> isUb R S x"
+  by (simp add: isUb_def)
+
+lemma isLub_le_isUb: "isLub R S x \<Longrightarrow> isUb R S y \<Longrightarrow> x \<le> y"
+  unfolding isLub_def by (blast intro!: leastPD3)
+
+lemma isLub_ubs: "isLub R S x \<Longrightarrow> x <=* ubs R S"
+  unfolding ubs_def isLub_def by (rule leastPD2)
+
+lemma isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::'a::linorder)"
+  apply (frule isLub_isUb)
+  apply (frule_tac x = y in isLub_isUb)
+  apply (blast intro!: order_antisym dest!: isLub_le_isUb)
+  done
+
+lemma isUb_UNIV_I: "(\<And>y. y \<in> S \<Longrightarrow> y \<le> u) \<Longrightarrow> isUb UNIV S u"
+  by (simp add: isUbI setleI)
+
+
+definition greatestP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a::ord \<Rightarrow> bool"
+  where "greatestP P x = (P x \<and> Collect P *<=  x)"
+
+definition isLb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
+  where "isLb R S x = (x <=* S \<and> x: R)"
+
+definition isGlb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
+  where "isGlb R S x = greatestP (isLb R S) x"
+
+definition lbs :: "'a set \<Rightarrow> 'a::ord set \<Rightarrow> 'a set"
+  where "lbs R S = Collect (isLb R S)"
+
+
+subsection {* Rules about the Operators @{term greatestP}, @{term isLb} and @{term isGlb} *}
+
+lemma greatestPD1: "greatestP P x \<Longrightarrow> P x"
+  by (simp add: greatestP_def)
+
+lemma greatestPD2: "greatestP P x \<Longrightarrow> Collect P *<= x"
+  by (simp add: greatestP_def)
+
+lemma greatestPD3: "greatestP P x \<Longrightarrow> y: Collect P \<Longrightarrow> x \<ge> y"
+  by (blast dest!: greatestPD2 setleD)
+
+lemma isGlbD1: "isGlb R S x \<Longrightarrow> x <=* S"
+  by (simp add: isGlb_def isLb_def greatestP_def)
+
+lemma isGlbD1a: "isGlb R S x \<Longrightarrow> x: R"
+  by (simp add: isGlb_def isLb_def greatestP_def)
+
+lemma isGlb_isLb: "isGlb R S x \<Longrightarrow> isLb R S x"
+  unfolding isLb_def by (blast dest: isGlbD1 isGlbD1a)
+
+lemma isGlbD2: "isGlb R S x \<Longrightarrow> y : S \<Longrightarrow> y \<ge> x"
+  by (blast dest!: isGlbD1 setgeD)
+
+lemma isGlbD3: "isGlb R S x \<Longrightarrow> greatestP (isLb R S) x"
+  by (simp add: isGlb_def)
+
+lemma isGlbI1: "greatestP (isLb R S) x \<Longrightarrow> isGlb R S x"
+  by (simp add: isGlb_def)
+
+lemma isGlbI2: "isLb R S x \<Longrightarrow> Collect (isLb R S) *<= x \<Longrightarrow> isGlb R S x"
+  by (simp add: isGlb_def greatestP_def)
+
+lemma isLbD: "isLb R S x \<Longrightarrow> y : S \<Longrightarrow> y \<ge> x"
+  by (simp add: isLb_def setge_def)
+
+lemma isLbD2: "isLb R S x \<Longrightarrow> x <=* S "
+  by (simp add: isLb_def)
+
+lemma isLbD2a: "isLb R S x \<Longrightarrow> x: R"
+  by (simp add: isLb_def)
+
+lemma isLbI: "x <=* S \<Longrightarrow> x: R \<Longrightarrow> isLb R S x"
+  by (simp add: isLb_def)
+
+lemma isGlb_le_isLb: "isGlb R S x \<Longrightarrow> isLb R S y \<Longrightarrow> x \<ge> y"
+  unfolding isGlb_def by (blast intro!: greatestPD3)
+
+lemma isGlb_ubs: "isGlb R S x \<Longrightarrow> lbs R S *<= x"
+  unfolding lbs_def isGlb_def by (rule greatestPD2)
+
+lemma isGlb_unique: "[| isGlb R S x; isGlb R S y |] ==> x = (y::'a::linorder)"
+  apply (frule isGlb_isLb)
+  apply (frule_tac x = y in isGlb_isLb)
+  apply (blast intro!: order_antisym dest!: isGlb_le_isLb)
+  done
+
+lemma bdd_above_setle: "bdd_above A \<longleftrightarrow> (\<exists>a. A *<= a)"
+  by (auto simp: bdd_above_def setle_def)
+
+lemma bdd_below_setge: "bdd_below A \<longleftrightarrow> (\<exists>a. a <=* A)"
+  by (auto simp: bdd_below_def setge_def)
+
+lemma isLub_cSup: 
+  "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> (\<exists>b. S *<= b) \<Longrightarrow> isLub UNIV S (Sup S)"
+  by  (auto simp add: isLub_def setle_def leastP_def isUb_def
+            intro!: setgeI cSup_upper cSup_least)
+
+lemma isGlb_cInf: 
+  "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> (\<exists>b. b <=* S) \<Longrightarrow> isGlb UNIV S (Inf S)"
+  by  (auto simp add: isGlb_def setge_def greatestP_def isLb_def
+            intro!: setleI cInf_lower cInf_greatest)
+
+lemma cSup_le: "(S::'a::conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> S *<= b \<Longrightarrow> Sup S \<le> b"
+  by (metis cSup_least setle_def)
+
+lemma cInf_ge: "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> b <=* S \<Longrightarrow> Inf S \<ge> b"
+  by (metis cInf_greatest setge_def)
+
+lemma cSup_bounds:
+  fixes S :: "'a :: conditionally_complete_lattice set"
+  shows "S \<noteq> {} \<Longrightarrow> a <=* S \<Longrightarrow> S *<= b \<Longrightarrow> a \<le> Sup S \<and> Sup S \<le> b"
+  using cSup_least[of S b] cSup_upper2[of _ S a]
+  by (auto simp: bdd_above_setle setge_def setle_def)
+
+lemma cSup_unique: "(S::'a :: {conditionally_complete_linorder, no_bot} set) *<= b \<Longrightarrow> (\<forall>b'<b. \<exists>x\<in>S. b' < x) \<Longrightarrow> Sup S = b"
+  by (rule cSup_eq) (auto simp: not_le[symmetric] setle_def)
+
+lemma cInf_unique: "b <=* (S::'a :: {conditionally_complete_linorder, no_top} set) \<Longrightarrow> (\<forall>b'>b. \<exists>x\<in>S. b' > x) \<Longrightarrow> Inf S = b"
+  by (rule cInf_eq) (auto simp: not_le[symmetric] setge_def)
+
+text{* Use completeness of reals (supremum property) to show that any bounded sequence has a least upper bound*}
+
+lemma reals_complete: "\<exists>X. X \<in> S \<Longrightarrow> \<exists>Y. isUb (UNIV::real set) S Y \<Longrightarrow> \<exists>t. isLub (UNIV :: real set) S t"
+  by (intro exI[of _ "Sup S"] isLub_cSup) (auto simp: setle_def isUb_def intro!: cSup_upper)
+
+lemma Bseq_isUb: "\<And>X :: nat \<Rightarrow> real. Bseq X \<Longrightarrow> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U"
+  by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff)
+
+lemma Bseq_isLub: "\<And>X :: nat \<Rightarrow> real. Bseq X \<Longrightarrow> \<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U"
+  by (blast intro: reals_complete Bseq_isUb)
+
+lemma isLub_mono_imp_LIMSEQ:
+  fixes X :: "nat \<Rightarrow> real"
+  assumes u: "isLub UNIV {x. \<exists>n. X n = x} u" (* FIXME: use 'range X' *)
+  assumes X: "\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n"
+  shows "X ----> u"
+proof -
+  have "X ----> (SUP i. X i)"
+    using u[THEN isLubD1] X
+    by (intro LIMSEQ_incseq_SUP) (auto simp: incseq_def image_def eq_commute bdd_above_setle)
+  also have "(SUP i. X i) = u"
+    using isLub_cSup[of "range X"] u[THEN isLubD1]
+    by (intro isLub_unique[OF _ u]) (auto simp add: SUP_def image_def eq_commute)
+  finally show ?thesis .
+qed
+
+lemmas real_isGlb_unique = isGlb_unique[where 'a=real]
+
+lemma real_le_inf_subset: "t \<noteq> {} \<Longrightarrow> t \<subseteq> s \<Longrightarrow> \<exists>b. b <=* s \<Longrightarrow> Inf s \<le> Inf (t::real set)"
+  by (rule cInf_superset_mono) (auto simp: bdd_below_setge)
+
+lemma real_ge_sup_subset: "t \<noteq> {} \<Longrightarrow> t \<subseteq> s \<Longrightarrow> \<exists>b. s *<= b \<Longrightarrow> Sup s \<ge> Sup (t::real set)"
+  by (rule cSup_subset_mono) (auto simp: bdd_above_setle)
+
+end
--- a/src/HOL/Library/Multiset.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Multiset.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1533,10 +1533,10 @@
   qed
 qed
 
-lemma all_accessible: "wf r ==> \<forall>M. M \<in> acc (mult1 r)"
+lemma all_accessible: "wf r ==> \<forall>M. M \<in> Wellfounded.acc (mult1 r)"
 proof
   let ?R = "mult1 r"
-  let ?W = "acc ?R"
+  let ?W = "Wellfounded.acc ?R"
   {
     fix M M0 a
     assume M0: "M0 \<in> ?W"
--- a/src/HOL/Library/Polynomial.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Polynomial.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -667,7 +667,7 @@
   show "- p + p = 0"
     by (simp add: poly_eq_iff)
   show "p - q = p + - q"
-    by (simp add: poly_eq_iff diff_minus)
+    by (simp add: poly_eq_iff)
 qed
 
 end
@@ -714,15 +714,15 @@
 
 lemma degree_diff_le_max: "degree (p - q) \<le> max (degree p) (degree q)"
   using degree_add_le [where p=p and q="-q"]
-  by (simp add: diff_minus)
+  by simp
 
 lemma degree_diff_le:
   "\<lbrakk>degree p \<le> n; degree q \<le> n\<rbrakk> \<Longrightarrow> degree (p - q) \<le> n"
-  by (simp add: diff_minus degree_add_le)
+  using degree_add_le [of p n "- q"] by simp
 
 lemma degree_diff_less:
   "\<lbrakk>degree p < n; degree q < n\<rbrakk> \<Longrightarrow> degree (p - q) < n"
-  by (simp add: diff_minus degree_add_less)
+  using degree_add_less [of p n "- q"] by simp
 
 lemma add_monom: "monom a n + monom b n = monom (a + b) n"
   by (rule poly_eqI) simp
@@ -793,7 +793,7 @@
 lemma poly_diff [simp]:
   fixes x :: "'a::comm_ring"
   shows "poly (p - q) x = poly p x - poly q x"
-  by (simp add: diff_minus)
+  using poly_add [of p "- q" x] by simp
 
 lemma poly_setsum: "poly (\<Sum>k\<in>A. p k) x = (\<Sum>k\<in>A. poly (p k) x)"
   by (induct A rule: infinite_finite_induct) simp_all
--- a/src/HOL/Library/Product_plus.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Product_plus.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -104,7 +104,7 @@
   (cancel_comm_monoid_add, cancel_comm_monoid_add) cancel_comm_monoid_add ..
 
 instance prod :: (group_add, group_add) group_add
-  by default (simp_all add: prod_eq_iff diff_minus)
+  by default (simp_all add: prod_eq_iff)
 
 instance prod :: (ab_group_add, ab_group_add) ab_group_add
   by default (simp_all add: prod_eq_iff)
--- a/src/HOL/Library/RBT_Set.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/RBT_Set.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -756,7 +756,8 @@
 declare Inf_Set_fold[folded Inf'_def, code]
 
 lemma INFI_Set_fold [code]:
-  "INFI (Set t) f = fold_keys (inf \<circ> f) t top"
+  fixes f :: "_ \<Rightarrow> 'a::complete_lattice"
+  shows "INFI (Set t) f = fold_keys (inf \<circ> f) t top"
 proof -
   have "comp_fun_commute ((inf :: 'a \<Rightarrow> 'a \<Rightarrow> 'a) \<circ> f)" 
     by default (auto simp add: fun_eq_iff ac_simps)
@@ -796,7 +797,8 @@
 declare Sup_Set_fold[folded Sup'_def, code]
 
 lemma SUPR_Set_fold [code]:
-  "SUPR (Set t) f = fold_keys (sup \<circ> f) t bot"
+  fixes f :: "_ \<Rightarrow> 'a::complete_lattice"
+  shows "SUPR (Set t) f = fold_keys (sup \<circ> f) t bot"
 proof -
   have "comp_fun_commute ((sup :: 'a \<Rightarrow> 'a \<Rightarrow> 'a) \<circ> f)" 
     by default (auto simp add: fun_eq_iff ac_simps)
--- a/src/HOL/Library/Set_Algebras.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/Set_Algebras.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -190,12 +190,12 @@
   done
 
 lemma set_plus_imp_minus: "(a::'a::ab_group_add) : b +o C ==> (a - b) : C"
-  by (auto simp add: elt_set_plus_def add_ac diff_minus)
+  by (auto simp add: elt_set_plus_def add_ac)
 
 lemma set_minus_imp_plus: "(a::'a::ab_group_add) - b : C ==> a : b +o C"
-  apply (auto simp add: elt_set_plus_def add_ac diff_minus)
+  apply (auto simp add: elt_set_plus_def add_ac)
   apply (subgoal_tac "a = (a + - b) + b")
-   apply (rule bexI, assumption, assumption)
+   apply (rule bexI, assumption)
   apply (auto simp add: add_ac)
   done
 
--- a/src/HOL/Library/Univ_Poly.thy	Mon Nov 11 17:34:44 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1053 +0,0 @@
-(*  Title:      HOL/Library/Univ_Poly.thy
-    Author:     Amine Chaieb
-*)
-
-header {* Univariate Polynomials *}
-
-theory Univ_Poly
-imports Main
-begin
-
-text{* Application of polynomial as a function. *}
-
-primrec (in semiring_0) poly :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a"
-where
-  poly_Nil:  "poly [] x = 0"
-| poly_Cons: "poly (h#t) x = h + x * poly t x"
-
-
-subsection{*Arithmetic Operations on Polynomials*}
-
-text{*addition*}
-
-primrec (in semiring_0) padd :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "+++" 65)
-where
-  padd_Nil:  "[] +++ l2 = l2"
-| padd_Cons: "(h#t) +++ l2 = (if l2 = [] then h#t else (h + hd l2)#(t +++ tl l2))"
-
-text{*Multiplication by a constant*}
-primrec (in semiring_0) cmult :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "%*" 70) where
-  cmult_Nil:  "c %* [] = []"
-| cmult_Cons: "c %* (h#t) = (c * h)#(c %* t)"
-
-text{*Multiplication by a polynomial*}
-primrec (in semiring_0) pmult :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "***" 70)
-where
-  pmult_Nil:  "[] *** l2 = []"
-| pmult_Cons: "(h#t) *** l2 = (if t = [] then h %* l2
-                              else (h %* l2) +++ ((0) # (t *** l2)))"
-
-text{*Repeated multiplication by a polynomial*}
-primrec (in semiring_0) mulexp :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a  list \<Rightarrow> 'a list" where
-  mulexp_zero:  "mulexp 0 p q = q"
-| mulexp_Suc:   "mulexp (Suc n) p q = p *** mulexp n p q"
-
-text{*Exponential*}
-primrec (in semiring_1) pexp :: "'a list \<Rightarrow> nat \<Rightarrow> 'a list"  (infixl "%^" 80) where
-  pexp_0:   "p %^ 0 = [1]"
-| pexp_Suc: "p %^ (Suc n) = p *** (p %^ n)"
-
-text{*Quotient related value of dividing a polynomial by x + a*}
-(* Useful for divisor properties in inductive proofs *)
-primrec (in field) "pquot" :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list"
-where
-  pquot_Nil:  "pquot [] a= []"
-| pquot_Cons: "pquot (h#t) a =
-    (if t = [] then [h] else (inverse(a) * (h - hd( pquot t a)))#(pquot t a))"
-
-text{*normalization of polynomials (remove extra 0 coeff)*}
-primrec (in semiring_0) pnormalize :: "'a list \<Rightarrow> 'a list" where
-  pnormalize_Nil:  "pnormalize [] = []"
-| pnormalize_Cons: "pnormalize (h#p) =
-    (if pnormalize p = [] then (if h = 0 then [] else [h]) else h # pnormalize p)"
-
-definition (in semiring_0) "pnormal p = ((pnormalize p = p) \<and> p \<noteq> [])"
-definition (in semiring_0) "nonconstant p = (pnormal p \<and> (\<forall>x. p \<noteq> [x]))"
-text{*Other definitions*}
-
-definition (in ring_1) poly_minus :: "'a list \<Rightarrow> 'a list" ("-- _" [80] 80)
-  where "-- p = (- 1) %* p"
-
-definition (in semiring_0) divides :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"  (infixl "divides" 70)
-  where "p1 divides p2 = (\<exists>q. poly p2 = poly(p1 *** q))"
-
-lemma (in semiring_0) dividesI:
-  "poly p2 = poly (p1 *** q) \<Longrightarrow> p1 divides p2"
-  by (auto simp add: divides_def)
-
-lemma (in semiring_0) dividesE:
-  assumes "p1 divides p2"
-  obtains q where "poly p2 = poly (p1 *** q)"
-  using assms by (auto simp add: divides_def)
-
-    --{*order of a polynomial*}
-definition (in ring_1) order :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where
-  "order a p = (SOME n. ([-a, 1] %^ n) divides p \<and> ~ (([-a, 1] %^ (Suc n)) divides p))"
-
-     --{*degree of a polynomial*}
-definition (in semiring_0) degree :: "'a list \<Rightarrow> nat"
-  where "degree p = length (pnormalize p) - 1"
-
-     --{*squarefree polynomials --- NB with respect to real roots only.*}
-definition (in ring_1) rsquarefree :: "'a list \<Rightarrow> bool"
-  where "rsquarefree p \<longleftrightarrow> poly p \<noteq> poly [] \<and> (\<forall>a. order a p = 0 \<or> order a p = 1)"
-
-context semiring_0
-begin
-
-lemma padd_Nil2[simp]: "p +++ [] = p"
-  by (induct p) auto
-
-lemma padd_Cons_Cons: "(h1 # p1) +++ (h2 # p2) = (h1 + h2) # (p1 +++ p2)"
-  by auto
-
-lemma pminus_Nil: "-- [] = []"
-  by (simp add: poly_minus_def)
-
-lemma pmult_singleton: "[h1] *** p1 = h1 %* p1" by simp
-
-end
-
-lemma (in semiring_1) poly_ident_mult[simp]: "1 %* t = t" by (induct t) auto
-
-lemma (in semiring_0) poly_simple_add_Cons[simp]: "[a] +++ ((0)#t) = (a#t)"
-  by simp
-
-text{*Handy general properties*}
-
-lemma (in comm_semiring_0) padd_commut: "b +++ a = a +++ b"
-proof (induct b arbitrary: a)
-  case Nil
-  thus ?case by auto
-next
-  case (Cons b bs a)
-  thus ?case by (cases a) (simp_all add: add_commute)
-qed
-
-lemma (in comm_semiring_0) padd_assoc: "\<forall>b c. (a +++ b) +++ c = a +++ (b +++ c)"
-  apply (induct a)
-  apply (simp, clarify)
-  apply (case_tac b, simp_all add: add_ac)
-  done
-
-lemma (in semiring_0) poly_cmult_distr: "a %* ( p +++ q) = (a %* p +++ a %* q)"
-  apply (induct p arbitrary: q)
-  apply simp
-  apply (case_tac q, simp_all add: distrib_left)
-  done
-
-lemma (in ring_1) pmult_by_x[simp]: "[0, 1] *** t = ((0)#t)"
-  apply (induct t)
-  apply simp
-  apply (auto simp add: padd_commut)
-  apply (case_tac t, auto)
-  done
-
-text{*properties of evaluation of polynomials.*}
-
-lemma (in semiring_0) poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x"
-proof(induct p1 arbitrary: p2)
-  case Nil
-  thus ?case by simp
-next
-  case (Cons a as p2)
-  thus ?case
-    by (cases p2) (simp_all  add: add_ac distrib_left)
-qed
-
-lemma (in comm_semiring_0) poly_cmult: "poly (c %* p) x = c * poly p x"
-  apply (induct p)
-  apply (case_tac [2] "x = zero")
-  apply (auto simp add: distrib_left mult_ac)
-  done
-
-lemma (in comm_semiring_0) poly_cmult_map: "poly (map (op * c) p) x = c*poly p x"
-  by (induct p) (auto simp add: distrib_left mult_ac)
-
-lemma (in comm_ring_1) poly_minus: "poly (-- p) x = - (poly p x)"
-  apply (simp add: poly_minus_def)
-  apply (auto simp add: poly_cmult)
-  done
-
-lemma (in comm_semiring_0) poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x"
-proof (induct p1 arbitrary: p2)
-  case Nil
-  thus ?case by simp
-next
-  case (Cons a as p2)
-  thus ?case by (cases as)
-    (simp_all add: poly_cmult poly_add distrib_right distrib_left mult_ac)
-qed
-
-class idom_char_0 = idom + ring_char_0
-
-lemma (in comm_ring_1) poly_exp: "poly (p %^ n) x = (poly p x) ^ n"
-  by (induct n) (auto simp add: poly_cmult poly_mult)
-
-text{*More Polynomial Evaluation Lemmas*}
-
-lemma (in semiring_0) poly_add_rzero[simp]: "poly (a +++ []) x = poly a x"
-  by simp
-
-lemma (in comm_semiring_0) poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x"
-  by (simp add: poly_mult mult_assoc)
-
-lemma (in semiring_0) poly_mult_Nil2[simp]: "poly (p *** []) x = 0"
-  by (induct p) auto
-
-lemma (in comm_semiring_1) poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x"
-  by (induct n) (auto simp add: poly_mult mult_assoc)
-
-subsection{*Key Property: if @{term "f(a) = 0"} then @{term "(x - a)"} divides
- @{term "p(x)"} *}
-
-lemma (in comm_ring_1) lemma_poly_linear_rem: "\<forall>h. \<exists>q r. h#t = [r] +++ [-a, 1] *** q"
-proof(induct t)
-  case Nil
-  { fix h have "[h] = [h] +++ [- a, 1] *** []" by simp }
-  thus ?case by blast
-next
-  case (Cons  x xs)
-  { fix h
-    from Cons.hyps[rule_format, of x]
-    obtain q r where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast
-    have "h#x#xs = [a*r + h] +++ [-a, 1] *** (r#q)"
-      using qr by (cases q) (simp_all add: algebra_simps)
-    hence "\<exists>q r. h#x#xs = [r] +++ [-a, 1] *** q" by blast}
-  thus ?case by blast
-qed
-
-lemma (in comm_ring_1) poly_linear_rem: "\<exists>q r. h#t = [r] +++ [-a, 1] *** q"
-  using lemma_poly_linear_rem [where t = t and a = a] by auto
-
-
-lemma (in comm_ring_1) poly_linear_divides: "(poly p a = 0) = ((p = []) | (\<exists>q. p = [-a, 1] *** q))"
-proof -
-  { assume p: "p = []" hence ?thesis by simp }
-  moreover
-  {
-    fix x xs assume p: "p = x#xs"
-    {
-      fix q assume "p = [-a, 1] *** q"
-      hence "poly p a = 0" by (simp add: poly_add poly_cmult)
-    }
-    moreover
-    { assume p0: "poly p a = 0"
-      from poly_linear_rem[of x xs a] obtain q r
-      where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast
-      have "r = 0" using p0 by (simp only: p qr poly_mult poly_add) simp
-      hence "\<exists>q. p = [- a, 1] *** q"
-        using p qr
-        apply -
-        apply (rule exI[where x=q])
-        apply auto
-        apply (cases q)
-        apply auto
-        done
-    }
-    ultimately have ?thesis using p by blast
-  }
-  ultimately show ?thesis by (cases p) auto
-qed
-
-lemma (in semiring_0) lemma_poly_length_mult[simp]: "\<forall>h k a. length (k %* p +++  (h # (a %* p))) = Suc (length p)"
-  by (induct p) auto
-
-lemma (in semiring_0) lemma_poly_length_mult2[simp]: "\<forall>h k. length (k %* p +++  (h # p)) = Suc (length p)"
-  by (induct p) auto
-
-lemma (in ring_1) poly_length_mult[simp]: "length([-a,1] *** q) = Suc (length q)"
-  by auto
-
-subsection{*Polynomial length*}
-
-lemma (in semiring_0) poly_cmult_length[simp]: "length (a %* p) = length p"
-  by (induct p) auto
-
-lemma (in semiring_0) poly_add_length: "length (p1 +++ p2) = max (length p1) (length p2)"
-  by (induct p1 arbitrary: p2) (simp_all, arith)
-
-lemma (in semiring_0) poly_root_mult_length[simp]: "length([a,b] *** p) = Suc (length p)"
-  by (simp add: poly_add_length)
-
-lemma (in idom) poly_mult_not_eq_poly_Nil[simp]:
-  "poly (p *** q) x \<noteq> poly [] x \<longleftrightarrow> poly p x \<noteq> poly [] x \<and> poly q x \<noteq> poly [] x"
-  by (auto simp add: poly_mult)
-
-lemma (in idom) poly_mult_eq_zero_disj: "poly (p *** q) x = 0 \<longleftrightarrow> poly p x = 0 \<or> poly q x = 0"
-  by (auto simp add: poly_mult)
-
-text{*Normalisation Properties*}
-
-lemma (in semiring_0) poly_normalized_nil: "(pnormalize p = []) --> (poly p x = 0)"
-  by (induct p) auto
-
-text{*A nontrivial polynomial of degree n has no more than n roots*}
-lemma (in idom) poly_roots_index_lemma:
-   assumes p: "poly p x \<noteq> poly [] x" and n: "length p = n"
-  shows "\<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)"
-  using p n
-proof (induct n arbitrary: p x)
-  case 0
-  thus ?case by simp
-next
-  case (Suc n p x)
-  {
-    assume C: "\<And>i. \<exists>x. poly p x = 0 \<and> (\<forall>m\<le>Suc n. x \<noteq> i m)"
-    from Suc.prems have p0: "poly p x \<noteq> 0" "p\<noteq> []" by auto
-    from p0(1)[unfolded poly_linear_divides[of p x]]
-    have "\<forall>q. p \<noteq> [- x, 1] *** q" by blast
-    from C obtain a where a: "poly p a = 0" by blast
-    from a[unfolded poly_linear_divides[of p a]] p0(2)
-    obtain q where q: "p = [-a, 1] *** q" by blast
-    have lg: "length q = n" using q Suc.prems(2) by simp
-    from q p0 have qx: "poly q x \<noteq> poly [] x"
-      by (auto simp add: poly_mult poly_add poly_cmult)
-    from Suc.hyps[OF qx lg] obtain i where
-      i: "\<forall>x. poly q x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)" by blast
-    let ?i = "\<lambda>m. if m = Suc n then a else i m"
-    from C[of ?i] obtain y where y: "poly p y = 0" "\<forall>m\<le> Suc n. y \<noteq> ?i m"
-      by blast
-    from y have "y = a \<or> poly q y = 0"
-      by (simp only: q poly_mult_eq_zero_disj poly_add) (simp add: algebra_simps)
-    with i[rule_format, of y] y(1) y(2) have False
-      apply auto
-      apply (erule_tac x = "m" in allE)
-      apply auto
-      done
-  }
-  thus ?case by blast
-qed
-
-
-lemma (in idom) poly_roots_index_length:
-  "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>i. \<forall>x. (poly p x = 0) \<longrightarrow> (\<exists>n. n \<le> length p \<and> x = i n)"
-  by (blast intro: poly_roots_index_lemma)
-
-lemma (in idom) poly_roots_finite_lemma1:
-  "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>N i. \<forall>x. (poly p x = 0) \<longrightarrow> (\<exists>n. (n::nat) < N \<and> x = i n)"
-  apply (drule poly_roots_index_length, safe)
-  apply (rule_tac x = "Suc (length p)" in exI)
-  apply (rule_tac x = i in exI)
-  apply (simp add: less_Suc_eq_le)
-  done
-
-lemma (in idom) idom_finite_lemma:
-  assumes P: "\<forall>x. P x --> (\<exists>n. n < length j \<and> x = j!n)"
-  shows "finite {x. P x}"
-proof -
-  let ?M = "{x. P x}"
-  let ?N = "set j"
-  have "?M \<subseteq> ?N" using P by auto
-  thus ?thesis using finite_subset by auto
-qed
-
-lemma (in idom) poly_roots_finite_lemma2:
-  "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> x \<in> set i"
-  apply (drule poly_roots_index_length, safe)
-  apply (rule_tac x="map (\<lambda>n. i n) [0 ..< Suc (length p)]" in exI)
-  apply (auto simp add: image_iff)
-  apply (erule_tac x="x" in allE, clarsimp)
-  apply (case_tac "n = length p")
-  apply (auto simp add: order_le_less)
-  done
-
-lemma (in ring_char_0) UNIV_ring_char_0_infinte: "\<not> (finite (UNIV:: 'a set))"
-proof
-  assume F: "finite (UNIV :: 'a set)"
-  have "finite (UNIV :: nat set)"
-  proof (rule finite_imageD)
-    have "of_nat ` UNIV \<subseteq> UNIV" by simp
-    then show "finite (of_nat ` UNIV :: 'a set)" using F by (rule finite_subset)
-    show "inj (of_nat :: nat \<Rightarrow> 'a)" by (simp add: inj_on_def)
-  qed
-  with infinite_UNIV_nat show False ..
-qed
-
-lemma (in idom_char_0) poly_roots_finite: "poly p \<noteq> poly [] \<longleftrightarrow> finite {x. poly p x = 0}"
-proof
-  assume H: "poly p \<noteq> poly []"
-  show "finite {x. poly p x = (0::'a)}"
-    using H
-    apply -
-    apply (erule contrapos_np, rule ext)
-    apply (rule ccontr)
-    apply (clarify dest!: poly_roots_finite_lemma2)
-    using finite_subset
-  proof -
-    fix x i
-    assume F: "\<not> finite {x. poly p x = (0\<Colon>'a)}"
-      and P: "\<forall>x. poly p x = (0\<Colon>'a) \<longrightarrow> x \<in> set i"
-    let ?M= "{x. poly p x = (0\<Colon>'a)}"
-    from P have "?M \<subseteq> set i" by auto
-    with finite_subset F show False by auto
-  qed
-next
-  assume F: "finite {x. poly p x = (0\<Colon>'a)}"
-  show "poly p \<noteq> poly []" using F UNIV_ring_char_0_infinte by auto
-qed
-
-text{*Entirety and Cancellation for polynomials*}
-
-lemma (in idom_char_0) poly_entire_lemma2:
-  assumes p0: "poly p \<noteq> poly []"
-    and q0: "poly q \<noteq> poly []"
-  shows "poly (p***q) \<noteq> poly []"
-proof -
-  let ?S = "\<lambda>p. {x. poly p x = 0}"
-  have "?S (p *** q) = ?S p \<union> ?S q" by (auto simp add: poly_mult)
-  with p0 q0 show ?thesis  unfolding poly_roots_finite by auto
-qed
-
-lemma (in idom_char_0) poly_entire:
-  "poly (p *** q) = poly [] \<longleftrightarrow> poly p = poly [] \<or> poly q = poly []"
-  using poly_entire_lemma2[of p q]
-  by (auto simp add: fun_eq_iff poly_mult)
-
-lemma (in idom_char_0) poly_entire_neg:
-  "poly (p *** q) \<noteq> poly [] \<longleftrightarrow> poly p \<noteq> poly [] \<and> poly q \<noteq> poly []"
-  by (simp add: poly_entire)
-
-lemma fun_eq: "f = g \<longleftrightarrow> (\<forall>x. f x = g x)"
-  by auto
-
-lemma (in comm_ring_1) poly_add_minus_zero_iff:
-  "poly (p +++ -- q) = poly [] \<longleftrightarrow> poly p = poly q"
-  by (auto simp add: algebra_simps poly_add poly_minus_def fun_eq poly_cmult)
-
-lemma (in comm_ring_1) poly_add_minus_mult_eq:
-  "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))"
-  by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult distrib_left)
-
-subclass (in idom_char_0) comm_ring_1 ..
-
-lemma (in idom_char_0) poly_mult_left_cancel:
-  "poly (p *** q) = poly (p *** r) \<longleftrightarrow> poly p = poly [] \<or> poly q = poly r"
-proof -
-  have "poly (p *** q) = poly (p *** r) \<longleftrightarrow> poly (p *** q +++ -- (p *** r)) = poly []"
-    by (simp only: poly_add_minus_zero_iff)
-  also have "\<dots> \<longleftrightarrow> poly p = poly [] \<or> poly q = poly r"
-    by (auto intro: simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff)
-  finally show ?thesis .
-qed
-
-lemma (in idom) poly_exp_eq_zero[simp]:
-  "poly (p %^ n) = poly [] \<longleftrightarrow> poly p = poly [] \<and> n \<noteq> 0"
-  apply (simp only: fun_eq add: HOL.all_simps [symmetric])
-  apply (rule arg_cong [where f = All])
-  apply (rule ext)
-  apply (induct n)
-  apply (auto simp add: poly_exp poly_mult)
-  done
-
-lemma (in comm_ring_1) poly_prime_eq_zero[simp]: "poly [a,1] \<noteq> poly []"
-  apply (simp add: fun_eq)
-  apply (rule_tac x = "minus one a" in exI)
-  apply (unfold diff_minus)
-  apply (subst add_commute)
-  apply (subst add_assoc)
-  apply simp
-  done
-
-lemma (in idom) poly_exp_prime_eq_zero: "poly ([a, 1] %^ n) \<noteq> poly []"
-  by auto
-
-text{*A more constructive notion of polynomials being trivial*}
-
-lemma (in idom_char_0) poly_zero_lemma': "poly (h # t) = poly [] \<Longrightarrow> h = 0 \<and> poly t = poly []"
-  apply (simp add: fun_eq)
-  apply (case_tac "h = zero")
-  apply (drule_tac [2] x = zero in spec, auto)
-  apply (cases "poly t = poly []", simp)
-proof -
-  fix x
-  assume H: "\<forall>x. x = (0\<Colon>'a) \<or> poly t x = (0\<Colon>'a)"
-    and pnz: "poly t \<noteq> poly []"
-  let ?S = "{x. poly t x = 0}"
-  from H have "\<forall>x. x \<noteq>0 \<longrightarrow> poly t x = 0" by blast
-  hence th: "?S \<supseteq> UNIV - {0}" by auto
-  from poly_roots_finite pnz have th': "finite ?S" by blast
-  from finite_subset[OF th th'] UNIV_ring_char_0_infinte show "poly t x = (0\<Colon>'a)"
-    by simp
-qed
-
-lemma (in idom_char_0) poly_zero: "(poly p = poly []) = list_all (%c. c = 0) p"
-  apply (induct p)
-  apply simp
-  apply (rule iffI)
-  apply (drule poly_zero_lemma', auto)
-  done
-
-lemma (in idom_char_0) poly_0: "list_all (\<lambda>c. c = 0) p \<Longrightarrow> poly p x = 0"
-  unfolding poly_zero[symmetric] by simp
-
-
-
-text{*Basics of divisibility.*}
-
-lemma (in idom) poly_primes:
-  "[a, 1] divides (p *** q) \<longleftrightarrow> [a, 1] divides p \<or> [a, 1] divides q"
-  apply (auto simp add: divides_def fun_eq poly_mult poly_add poly_cmult distrib_right [symmetric])
-  apply (drule_tac x = "uminus a" in spec)
-  apply (simp add: poly_linear_divides poly_add poly_cmult distrib_right [symmetric])
-  apply (cases "p = []")
-  apply (rule exI[where x="[]"])
-  apply simp
-  apply (cases "q = []")
-  apply (erule allE[where x="[]"], simp)
-
-  apply clarsimp
-  apply (cases "\<exists>q\<Colon>'a list. p = a %* q +++ ((0\<Colon>'a) # q)")
-  apply (clarsimp simp add: poly_add poly_cmult)
-  apply (rule_tac x="qa" in exI)
-  apply (simp add: distrib_right [symmetric])
-  apply clarsimp
-
-  apply (auto simp add: poly_linear_divides poly_add poly_cmult distrib_right [symmetric])
-  apply (rule_tac x = "pmult qa q" in exI)
-  apply (rule_tac [2] x = "pmult p qa" in exI)
-  apply (auto simp add: poly_add poly_mult poly_cmult mult_ac)
-  done
-
-lemma (in comm_semiring_1) poly_divides_refl[simp]: "p divides p"
-  apply (simp add: divides_def)
-  apply (rule_tac x = "[one]" in exI)
-  apply (auto simp add: poly_mult fun_eq)
-  done
-
-lemma (in comm_semiring_1) poly_divides_trans: "p divides q \<Longrightarrow> q divides r \<Longrightarrow> p divides r"
-  apply (simp add: divides_def, safe)
-  apply (rule_tac x = "pmult qa qaa" in exI)
-  apply (auto simp add: poly_mult fun_eq mult_assoc)
-  done
-
-lemma (in comm_semiring_1) poly_divides_exp: "m \<le> n \<Longrightarrow> (p %^ m) divides (p %^ n)"
-  apply (auto simp add: le_iff_add)
-  apply (induct_tac k)
-  apply (rule_tac [2] poly_divides_trans)
-  apply (auto simp add: divides_def)
-  apply (rule_tac x = p in exI)
-  apply (auto simp add: poly_mult fun_eq mult_ac)
-  done
-
-lemma (in comm_semiring_1) poly_exp_divides:
-  "(p %^ n) divides q \<Longrightarrow> m \<le> n \<Longrightarrow> (p %^ m) divides q"
-  by (blast intro: poly_divides_exp poly_divides_trans)
-
-lemma (in comm_semiring_0) poly_divides_add:
-  "p divides q \<Longrightarrow> p divides r \<Longrightarrow> p divides (q +++ r)"
-  apply (simp add: divides_def, auto)
-  apply (rule_tac x = "padd qa qaa" in exI)
-  apply (auto simp add: poly_add fun_eq poly_mult distrib_left)
-  done
-
-lemma (in comm_ring_1) poly_divides_diff:
-  "p divides q \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides r"
-  apply (simp add: divides_def, auto)
-  apply (rule_tac x = "padd qaa (poly_minus qa)" in exI)
-  apply (auto simp add: poly_add fun_eq poly_mult poly_minus algebra_simps)
-  done
-
-lemma (in comm_ring_1) poly_divides_diff2:
-  "p divides r \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides q"
-  apply (erule poly_divides_diff)
-  apply (auto simp add: poly_add fun_eq poly_mult divides_def add_ac)
-  done
-
-lemma (in semiring_0) poly_divides_zero: "poly p = poly [] \<Longrightarrow> q divides p"
-  apply (simp add: divides_def)
-  apply (rule exI[where x="[]"])
-  apply (auto simp add: fun_eq poly_mult)
-  done
-
-lemma (in semiring_0) poly_divides_zero2 [simp]: "q divides []"
-  apply (simp add: divides_def)
-  apply (rule_tac x = "[]" in exI)
-  apply (auto simp add: fun_eq)
-  done
-
-text{*At last, we can consider the order of a root.*}
-
-lemma (in idom_char_0) poly_order_exists_lemma:
-  assumes lp: "length p = d"
-    and p: "poly p \<noteq> poly []"
-  shows "\<exists>n q. p = mulexp n [-a, 1] q \<and> poly q a \<noteq> 0"
-  using lp p
-proof (induct d arbitrary: p)
-  case 0
-  thus ?case by simp
-next
-  case (Suc n p)
-  show ?case
-  proof (cases "poly p a = 0")
-    case True
-    from Suc.prems have h: "length p = Suc n" "poly p \<noteq> poly []" by auto
-    hence pN: "p \<noteq> []" by auto
-    from True[unfolded poly_linear_divides] pN obtain q where q: "p = [-a, 1] *** q"
-      by blast
-    from q h True have qh: "length q = n" "poly q \<noteq> poly []"
-      apply -
-      apply simp
-      apply (simp only: fun_eq)
-      apply (rule ccontr)
-      apply (simp add: fun_eq poly_add poly_cmult)
-      done
-    from Suc.hyps[OF qh] obtain m r where mr: "q = mulexp m [-a,1] r" "poly r a \<noteq> 0"
-      by blast
-    from mr q have "p = mulexp (Suc m) [-a,1] r \<and> poly r a \<noteq> 0" by simp
-    then show ?thesis by blast
-  next
-    case False
-    then show ?thesis
-      using Suc.prems
-      apply simp
-      apply (rule exI[where x="0::nat"])
-      apply simp
-      done
-  qed
-qed
-
-
-lemma (in comm_semiring_1) poly_mulexp: "poly (mulexp n p q) x = (poly p x) ^ n * poly q x"
-  by (induct n) (auto simp add: poly_mult mult_ac)
-
-lemma (in comm_semiring_1) divides_left_mult:
-  assumes d:"(p***q) divides r" shows "p divides r \<and> q divides r"
-proof-
-  from d obtain t where r:"poly r = poly (p***q *** t)"
-    unfolding divides_def by blast
-  hence "poly r = poly (p *** (q *** t))"
-    "poly r = poly (q *** (p***t))" by(auto simp add: fun_eq poly_mult mult_ac)
-  thus ?thesis unfolding divides_def by blast
-qed
-
-
-(* FIXME: Tidy up *)
-
-lemma (in semiring_1) zero_power_iff: "0 ^ n = (if n = 0 then 1 else 0)"
-  by (induct n) simp_all
-
-lemma (in idom_char_0) poly_order_exists:
-  assumes "length p = d" and "poly p \<noteq> poly []"
-  shows "\<exists>n. [- a, 1] %^ n divides p \<and> \<not> [- a, 1] %^ Suc n divides p"
-proof -
-  from assms have "\<exists>n q. p = mulexp n [- a, 1] q \<and> poly q a \<noteq> 0"
-    by (rule poly_order_exists_lemma)
-  then obtain n q where p: "p = mulexp n [- a, 1] q" and "poly q a \<noteq> 0" by blast
-  have "[- a, 1] %^ n divides mulexp n [- a, 1] q"
-  proof (rule dividesI)
-    show "poly (mulexp n [- a, 1] q) = poly ([- a, 1] %^ n *** q)"
-      by (induct n) (simp_all add: poly_add poly_cmult poly_mult distrib_left mult_ac)
-  qed
-  moreover have "\<not> [- a, 1] %^ Suc n divides mulexp n [- a, 1] q"
-  proof
-    assume "[- a, 1] %^ Suc n divides mulexp n [- a, 1] q"
-    then obtain m where "poly (mulexp n [- a, 1] q) = poly ([- a, 1] %^ Suc n *** m)"
-      by (rule dividesE)
-    moreover have "poly (mulexp n [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc n *** m)"
-    proof (induct n)
-      case 0 show ?case
-      proof (rule ccontr)
-        assume "\<not> poly (mulexp 0 [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc 0 *** m)"
-        then have "poly q a = 0"
-          by (simp add: poly_add poly_cmult)
-        with `poly q a \<noteq> 0` show False by simp
-      qed
-    next
-      case (Suc n) show ?case
-        by (rule pexp_Suc [THEN ssubst], rule ccontr)
-          (simp add: poly_mult_left_cancel poly_mult_assoc Suc del: pmult_Cons pexp_Suc)
-    qed
-    ultimately show False by simp
-  qed
-  ultimately show ?thesis by (auto simp add: p)
-qed
-
-lemma (in semiring_1) poly_one_divides[simp]: "[1] divides p"
-  by (auto simp add: divides_def)
-
-lemma (in idom_char_0) poly_order:
-  "poly p \<noteq> poly [] \<Longrightarrow> \<exists>!n. ([-a, 1] %^ n) divides p \<and> \<not> (([-a, 1] %^ Suc n) divides p)"
-  apply (auto intro: poly_order_exists simp add: less_linear simp del: pmult_Cons pexp_Suc)
-  apply (cut_tac x = y and y = n in less_linear)
-  apply (drule_tac m = n in poly_exp_divides)
-  apply (auto dest: Suc_le_eq [THEN iffD2, THEN [2] poly_exp_divides]
-              simp del: pmult_Cons pexp_Suc)
-  done
-
-text{*Order*}
-
-lemma some1_equalityD: "n = (SOME n. P n) \<Longrightarrow> \<exists>!n. P n \<Longrightarrow> P n"
-  by (blast intro: someI2)
-
-lemma (in idom_char_0) order:
-      "(([-a, 1] %^ n) divides p \<and>
-        ~(([-a, 1] %^ (Suc n)) divides p)) =
-        ((n = order a p) \<and> ~(poly p = poly []))"
-  apply (unfold order_def)
-  apply (rule iffI)
-  apply (blast dest: poly_divides_zero intro!: some1_equality [symmetric] poly_order)
-  apply (blast intro!: poly_order [THEN [2] some1_equalityD])
-  done
-
-lemma (in idom_char_0) order2:
-  "poly p \<noteq> poly [] \<Longrightarrow>
-    ([-a, 1] %^ (order a p)) divides p \<and> \<not> (([-a, 1] %^ (Suc (order a p))) divides p)"
-  by (simp add: order del: pexp_Suc)
-
-lemma (in idom_char_0) order_unique:
-  "poly p \<noteq> poly [] \<Longrightarrow> ([-a, 1] %^ n) divides p \<Longrightarrow> ~(([-a, 1] %^ (Suc n)) divides p) \<Longrightarrow>
-    n = order a p"
-  using order [of a n p] by auto
-
-lemma (in idom_char_0) order_unique_lemma:
-  "poly p \<noteq> poly [] \<and> ([-a, 1] %^ n) divides p \<and> ~(([-a, 1] %^ (Suc n)) divides p) \<Longrightarrow>
-    n = order a p"
-  by (blast intro: order_unique)
-
-lemma (in ring_1) order_poly: "poly p = poly q \<Longrightarrow> order a p = order a q"
-  by (auto simp add: fun_eq divides_def poly_mult order_def)
-
-lemma (in semiring_1) pexp_one[simp]: "p %^ (Suc 0) = p"
-  by (induct "p") auto
-
-lemma (in comm_ring_1) lemma_order_root:
-  "0 < n \<and> [- a, 1] %^ n divides p \<and> ~ [- a, 1] %^ (Suc n) divides p \<Longrightarrow> poly p a = 0"
-  by (induct n arbitrary: a p) (auto simp add: divides_def poly_mult simp del: pmult_Cons)
-
-lemma (in idom_char_0) order_root:
-  "poly p a = 0 \<longleftrightarrow> poly p = poly [] \<or> order a p \<noteq> 0"
-  apply (cases "poly p = poly []")
-  apply auto
-  apply (simp add: poly_linear_divides del: pmult_Cons, safe)
-  apply (drule_tac [!] a = a in order2)
-  apply (rule ccontr)
-  apply (simp add: divides_def poly_mult fun_eq del: pmult_Cons, blast)
-  using neq0_conv
-  apply (blast intro: lemma_order_root)
-  done
-
-lemma (in idom_char_0) order_divides:
-  "([-a, 1] %^ n) divides p \<longleftrightarrow> poly p = poly [] \<or> n \<le> order a p"
-  apply (cases "poly p = poly []")
-  apply auto
-  apply (simp add: divides_def fun_eq poly_mult)
-  apply (rule_tac x = "[]" in exI)
-  apply (auto dest!: order2 [where a=a] intro: poly_exp_divides simp del: pexp_Suc)
-  done
-
-lemma (in idom_char_0) order_decomp:
-  "poly p \<noteq> poly [] \<Longrightarrow> \<exists>q. poly p = poly (([-a, 1] %^ (order a p)) *** q) \<and> ~([-a, 1] divides q)"
-  apply (unfold divides_def)
-  apply (drule order2 [where a = a])
-  apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe)
-  apply (rule_tac x = q in exI, safe)
-  apply (drule_tac x = qa in spec)
-  apply (auto simp add: poly_mult fun_eq poly_exp mult_ac simp del: pmult_Cons)
-  done
-
-text{*Important composition properties of orders.*}
-lemma order_mult:
-  "poly (p *** q) \<noteq> poly [] \<Longrightarrow>
-    order a (p *** q) = order a p + order (a::'a::{idom_char_0}) q"
-  apply (cut_tac a = a and p = "p *** q" and n = "order a p + order a q" in order)
-  apply (auto simp add: poly_entire simp del: pmult_Cons)
-  apply (drule_tac a = a in order2)+
-  apply safe
-  apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe)
-  apply (rule_tac x = "qa *** qaa" in exI)
-  apply (simp add: poly_mult mult_ac del: pmult_Cons)
-  apply (drule_tac a = a in order_decomp)+
-  apply safe
-  apply (subgoal_tac "[-a,1] divides (qa *** qaa) ")
-  apply (simp add: poly_primes del: pmult_Cons)
-  apply (auto simp add: divides_def simp del: pmult_Cons)
-  apply (rule_tac x = qb in exI)
-  apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) = poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))")
-  apply (drule poly_mult_left_cancel [THEN iffD1], force)
-  apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) = poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ")
-  apply (drule poly_mult_left_cancel [THEN iffD1], force)
-  apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
-  done
-
-lemma (in idom_char_0) order_mult:
-  assumes "poly (p *** q) \<noteq> poly []"
-  shows "order a (p *** q) = order a p + order a q"
-  using assms
-  apply (cut_tac a = a and p = "pmult p q" and n = "order a p + order a q" in order)
-  apply (auto simp add: poly_entire simp del: pmult_Cons)
-  apply (drule_tac a = a in order2)+
-  apply safe
-  apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe)
-  apply (rule_tac x = "pmult qa qaa" in exI)
-  apply (simp add: poly_mult mult_ac del: pmult_Cons)
-  apply (drule_tac a = a in order_decomp)+
-  apply safe
-  apply (subgoal_tac "[uminus a, one] divides pmult qa qaa")
-  apply (simp add: poly_primes del: pmult_Cons)
-  apply (auto simp add: divides_def simp del: pmult_Cons)
-  apply (rule_tac x = qb in exI)
-  apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa)) =
-    poly (pmult (pexp [uminus a, one] (?order a p)) (pmult [uminus a, one] qb))")
-  apply (drule poly_mult_left_cancel [THEN iffD1], force)
-  apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a q))
-      (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa))) =
-    poly (pmult (pexp [uminus a, one] (order a q))
-      (pmult (pexp [uminus a, one] (order a p)) (pmult [uminus a, one] qb)))")
-  apply (drule poly_mult_left_cancel [THEN iffD1], force)
-  apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
-  done
-
-lemma (in idom_char_0) order_root2: "poly p \<noteq> poly [] \<Longrightarrow> poly p a = 0 \<longleftrightarrow> order a p \<noteq> 0"
-  by (rule order_root [THEN ssubst]) auto
-
-lemma (in semiring_1) pmult_one[simp]: "[1] *** p = p" by auto
-
-lemma (in semiring_0) poly_Nil_zero: "poly [] = poly [0]"
-  by (simp add: fun_eq)
-
-lemma (in idom_char_0) rsquarefree_decomp:
-  "rsquarefree p \<Longrightarrow> poly p a = 0 \<Longrightarrow>
-    \<exists>q. poly p = poly ([-a, 1] *** q) \<and> poly q a \<noteq> 0"
-  apply (simp add: rsquarefree_def, safe)
-  apply (frule_tac a = a in order_decomp)
-  apply (drule_tac x = a in spec)
-  apply (drule_tac a = a in order_root2 [symmetric])
-  apply (auto simp del: pmult_Cons)
-  apply (rule_tac x = q in exI, safe)
-  apply (simp add: poly_mult fun_eq)
-  apply (drule_tac p1 = q in poly_linear_divides [THEN iffD1])
-  apply (simp add: divides_def del: pmult_Cons, safe)
-  apply (drule_tac x = "[]" in spec)
-  apply (auto simp add: fun_eq)
-  done
-
-
-text{*Normalization of a polynomial.*}
-
-lemma (in semiring_0) poly_normalize[simp]: "poly (pnormalize p) = poly p"
-  by (induct p) (auto simp add: fun_eq)
-
-text{*The degree of a polynomial.*}
-
-lemma (in semiring_0) lemma_degree_zero: "list_all (%c. c = 0) p \<longleftrightarrow> pnormalize p = []"
-  by (induct p) auto
-
-lemma (in idom_char_0) degree_zero:
-  assumes "poly p = poly []"
-  shows "degree p = 0"
-  using assms
-  by (cases "pnormalize p = []") (auto simp add: degree_def poly_zero lemma_degree_zero)
-
-lemma (in semiring_0) pnormalize_sing: "(pnormalize [x] = [x]) \<longleftrightarrow> x \<noteq> 0"
-  by simp
-
-lemma (in semiring_0) pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> (pnormalize [x, y] = [x, y])"
-  by simp
-
-lemma (in semiring_0) pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c#p)"
-  unfolding pnormal_def by simp
-
-lemma (in semiring_0) pnormal_tail: "p\<noteq>[] \<Longrightarrow> pnormal (c#p) \<Longrightarrow> pnormal p"
-  unfolding pnormal_def by(auto split: split_if_asm)
-
-
-lemma (in semiring_0) pnormal_last_nonzero: "pnormal p \<Longrightarrow> last p \<noteq> 0"
-  by (induct p) (simp_all add: pnormal_def split: split_if_asm)
-
-lemma (in semiring_0) pnormal_length: "pnormal p \<Longrightarrow> 0 < length p"
-  unfolding pnormal_def length_greater_0_conv by blast
-
-lemma (in semiring_0) pnormal_last_length: "0 < length p \<Longrightarrow> last p \<noteq> 0 \<Longrightarrow> pnormal p"
-  by (induct p) (auto simp: pnormal_def  split: split_if_asm)
-
-
-lemma (in semiring_0) pnormal_id: "pnormal p \<longleftrightarrow> 0 < length p \<and> last p \<noteq> 0"
-  using pnormal_last_length pnormal_length pnormal_last_nonzero by blast
-
-lemma (in idom_char_0) poly_Cons_eq:
-  "poly (c # cs) = poly (d # ds) \<longleftrightarrow> c = d \<and> poly cs = poly ds"
-  (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  assume eq: ?lhs
-  hence "\<And>x. poly ((c#cs) +++ -- (d#ds)) x = 0"
-    by (simp only: poly_minus poly_add algebra_simps) simp
-  hence "poly ((c#cs) +++ -- (d#ds)) = poly []" by(simp add: fun_eq_iff)
-  hence "c = d \<and> list_all (\<lambda>x. x=0) ((cs +++ -- ds))"
-    unfolding poly_zero by (simp add: poly_minus_def algebra_simps)
-  hence "c = d \<and> (\<forall>x. poly (cs +++ -- ds) x = 0)"
-    unfolding poly_zero[symmetric] by simp
-  then show ?rhs by (simp add: poly_minus poly_add algebra_simps fun_eq_iff)
-next
-  assume ?rhs
-  then show ?lhs by(simp add:fun_eq_iff)
-qed
-
-lemma (in idom_char_0) pnormalize_unique: "poly p = poly q \<Longrightarrow> pnormalize p = pnormalize q"
-proof (induct q arbitrary: p)
-  case Nil
-  thus ?case by (simp only: poly_zero lemma_degree_zero) simp
-next
-  case (Cons c cs p)
-  thus ?case
-  proof (induct p)
-    case Nil
-    hence "poly [] = poly (c#cs)" by blast
-    then have "poly (c#cs) = poly [] " by simp
-    thus ?case by (simp only: poly_zero lemma_degree_zero) simp
-  next
-    case (Cons d ds)
-    hence eq: "poly (d # ds) = poly (c # cs)" by blast
-    hence eq': "\<And>x. poly (d # ds) x = poly (c # cs) x" by simp
-    hence "poly (d # ds) 0 = poly (c # cs) 0" by blast
-    hence dc: "d = c" by auto
-    with eq have "poly ds = poly cs"
-      unfolding  poly_Cons_eq by simp
-    with Cons.prems have "pnormalize ds = pnormalize cs" by blast
-    with dc show ?case by simp
-  qed
-qed
-
-lemma (in idom_char_0) degree_unique:
-  assumes pq: "poly p = poly q"
-  shows "degree p = degree q"
-  using pnormalize_unique[OF pq] unfolding degree_def by simp
-
-lemma (in semiring_0) pnormalize_length:
-  "length (pnormalize p) \<le> length p" by (induct p) auto
-
-lemma (in semiring_0) last_linear_mul_lemma:
-  "last ((a %* p) +++ (x#(b %* p))) = (if p = [] then x else b * last p)"
-  apply (induct p arbitrary: a x b)
-  apply auto
-  apply (subgoal_tac "padd (cmult aa p) (times b a # cmult b p) \<noteq> []")
-  apply simp
-  apply (induct_tac p)
-  apply auto
-  done
-
-lemma (in semiring_1) last_linear_mul:
-  assumes p: "p \<noteq> []"
-  shows "last ([a,1] *** p) = last p"
-proof -
-  from p obtain c cs where cs: "p = c#cs" by (cases p) auto
-  from cs have eq: "[a,1] *** p = (a %* (c#cs)) +++ (0#(1 %* (c#cs)))"
-    by (simp add: poly_cmult_distr)
-  show ?thesis using cs
-    unfolding eq last_linear_mul_lemma by simp
-qed
-
-lemma (in semiring_0) pnormalize_eq: "last p \<noteq> 0 \<Longrightarrow> pnormalize p = p"
-  by (induct p) (auto split: split_if_asm)
-
-lemma (in semiring_0) last_pnormalize: "pnormalize p \<noteq> [] \<Longrightarrow> last (pnormalize p) \<noteq> 0"
-  by (induct p) auto
-
-lemma (in semiring_0) pnormal_degree: "last p \<noteq> 0 \<Longrightarrow> degree p = length p - 1"
-  using pnormalize_eq[of p] unfolding degree_def by simp
-
-lemma (in semiring_0) poly_Nil_ext: "poly [] = (\<lambda>x. 0)"
-  by (rule ext) simp
-
-lemma (in idom_char_0) linear_mul_degree:
-  assumes p: "poly p \<noteq> poly []"
-  shows "degree ([a,1] *** p) = degree p + 1"
-proof -
-  from p have pnz: "pnormalize p \<noteq> []"
-    unfolding poly_zero lemma_degree_zero .
-
-  from last_linear_mul[OF pnz, of a] last_pnormalize[OF pnz]
-  have l0: "last ([a, 1] *** pnormalize p) \<noteq> 0" by simp
-  from last_pnormalize[OF pnz] last_linear_mul[OF pnz, of a]
-    pnormal_degree[OF l0] pnormal_degree[OF last_pnormalize[OF pnz]] pnz
-
-  have th: "degree ([a,1] *** pnormalize p) = degree (pnormalize p) + 1"
-    by simp
-
-  have eqs: "poly ([a,1] *** pnormalize p) = poly ([a,1] *** p)"
-    by (rule ext) (simp add: poly_mult poly_add poly_cmult)
-  from degree_unique[OF eqs] th
-  show ?thesis by (simp add: degree_unique[OF poly_normalize])
-qed
-
-lemma (in idom_char_0) linear_pow_mul_degree:
-  "degree([a,1] %^n *** p) = (if poly p = poly [] then 0 else degree p + n)"
-proof (induct n arbitrary: a p)
-  case (0 a p)
-  show ?case
-  proof (cases "poly p = poly []")
-    case True
-    then show ?thesis
-      using degree_unique[OF True] by (simp add: degree_def)
-  next
-    case False
-    then show ?thesis by (auto simp add: poly_Nil_ext)
-  qed
-next
-  case (Suc n a p)
-  have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1] %^ n *** ([a,1] *** p))"
-    apply (rule ext)
-    apply (simp add: poly_mult poly_add poly_cmult)
-    apply (simp add: mult_ac add_ac distrib_left)
-    done
-  note deq = degree_unique[OF eq]
-  show ?case
-  proof (cases "poly p = poly []")
-    case True
-    with eq have eq': "poly ([a,1] %^(Suc n) *** p) = poly []"
-      apply -
-      apply (rule ext)
-      apply (simp add: poly_mult poly_cmult poly_add)
-      done
-    from degree_unique[OF eq'] True show ?thesis
-      by (simp add: degree_def)
-  next
-    case False
-    then have ap: "poly ([a,1] *** p) \<noteq> poly []"
-      using poly_mult_not_eq_poly_Nil unfolding poly_entire by auto
-    have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1]%^n *** ([a,1] *** p))"
-      by (rule ext, simp add: poly_mult poly_add poly_exp poly_cmult algebra_simps)
-    from ap have ap': "(poly ([a,1] *** p) = poly []) = False"
-      by blast
-    have th0: "degree ([a,1]%^n *** ([a,1] *** p)) = degree ([a,1] *** p) + n"
-      apply (simp only: Suc.hyps[of a "pmult [a,one] p"] ap')
-      apply simp
-      done
-    from degree_unique[OF eq] ap False th0 linear_mul_degree[OF False, of a]
-    show ?thesis by (auto simp del: poly.simps)
-  qed
-qed
-
-lemma (in idom_char_0) order_degree:
-  assumes p0: "poly p \<noteq> poly []"
-  shows "order a p \<le> degree p"
-proof -
-  from order2[OF p0, unfolded divides_def]
-  obtain q where q: "poly p = poly ([- a, 1]%^ (order a p) *** q)" by blast
-  {
-    assume "poly q = poly []"
-    with q p0 have False by (simp add: poly_mult poly_entire)
-  }
-  with degree_unique[OF q, unfolded linear_pow_mul_degree] show ?thesis
-    by auto
-qed
-
-text{*Tidier versions of finiteness of roots.*}
-
-lemma (in idom_char_0) poly_roots_finite_set:
-  "poly p \<noteq> poly [] \<Longrightarrow> finite {x. poly p x = 0}"
-  unfolding poly_roots_finite .
-
-text{*bound for polynomial.*}
-
-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 (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)
-  done
-
-lemma (in semiring_0) poly_Sing: "poly [c] x = c" by simp
-
-end
--- a/src/HOL/Library/While_Combinator.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Library/While_Combinator.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -307,37 +307,44 @@
 by Nipkow (the theories are in the AFP entry Flyspeck by Nipkow)
 and the AFP article Executable Transitive Closures by René Thiemann. *}
 
-definition rtrancl_while :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a list) \<Rightarrow> 'a
-  \<Rightarrow> ('a list * 'a set) option"
-where "rtrancl_while p f x =
-  while_option (%(ws,_). ws \<noteq> [] \<and> p(hd ws))
-    ((%(ws,Z).
-     let x = hd ws; new = filter (\<lambda>y. y \<notin> Z) (f x)
-     in (new @ tl ws, set new \<union> Z)))
-    ([x],{x})"
+context
+fixes p :: "'a \<Rightarrow> bool"
+and f :: "'a \<Rightarrow> 'a list"
+and x :: 'a
+begin
+
+fun rtrancl_while_test :: "'a list \<times> 'a set \<Rightarrow> bool"
+where "rtrancl_while_test (ws,_) = (ws \<noteq> [] \<and> p(hd ws))"
+
+fun rtrancl_while_step :: "'a list \<times> 'a set \<Rightarrow> 'a list \<times> 'a set"
+where "rtrancl_while_step (ws, Z) =
+  (let x = hd ws; new = remdups (filter (\<lambda>y. y \<notin> Z) (f x))
+  in (new @ tl ws, set new \<union> Z))"
 
-lemma rtrancl_while_Some: assumes "rtrancl_while p f x = Some(ws,Z)"
+definition rtrancl_while :: "('a list * 'a set) option"
+where "rtrancl_while = while_option rtrancl_while_test rtrancl_while_step ([x],{x})"
+
+fun rtrancl_while_invariant :: "'a list \<times> 'a set \<Rightarrow> bool"
+where "rtrancl_while_invariant (ws, Z) =
+   (x \<in> Z \<and> set ws \<subseteq> Z \<and> distinct ws \<and> {(x,y). y \<in> set(f x)} `` (Z - set ws) \<subseteq> Z \<and>
+    Z \<subseteq> {(x,y). y \<in> set(f x)}^* `` {x} \<and> (\<forall>z\<in>Z - set ws. p z))"
+
+lemma rtrancl_while_invariant: 
+  assumes inv: "rtrancl_while_invariant st" and test: "rtrancl_while_test st"
+  shows   "rtrancl_while_invariant (rtrancl_while_step st)"
+proof (cases st)
+  fix ws Z assume st: "st = (ws, Z)"
+  with test obtain h t where "ws = h # t" "p h" by (cases ws) auto
+  with inv st show ?thesis by (auto intro: rtrancl.rtrancl_into_rtrancl)
+qed
+
+lemma rtrancl_while_Some: assumes "rtrancl_while = Some(ws,Z)"
 shows "if ws = []
        then Z = {(x,y). y \<in> set(f x)}^* `` {x} \<and> (\<forall>z\<in>Z. p z)
        else \<not>p(hd ws) \<and> hd ws \<in> {(x,y). y \<in> set(f x)}^* `` {x}"
-proof-
-  let ?test = "(%(ws,_). ws \<noteq> [] \<and> p(hd ws))"
-  let ?step = "(%(ws,Z).
-     let x = hd ws; new = filter (\<lambda>y. y \<notin> Z) (f x)
-     in (new @ tl ws, set new \<union> Z))"
-  let ?R = "{(x,y). y \<in> set(f x)}"
-  let ?Inv = "%(ws,Z). x \<in> Z \<and> set ws \<subseteq> Z \<and> ?R `` (Z - set ws) \<subseteq> Z \<and>
-                       Z \<subseteq> ?R^* `` {x} \<and> (\<forall>z\<in>Z - set ws. p z)"
-  { fix ws Z assume 1: "?Inv(ws,Z)" and 2: "?test(ws,Z)"
-    from 2 obtain v vs where [simp]: "ws = v#vs" by (auto simp: neq_Nil_conv)
-    have "?Inv(?step (ws,Z))" using 1 2
-      by (auto intro: rtrancl.rtrancl_into_rtrancl)
-  } note inv = this
-  hence "!!p. ?Inv p \<Longrightarrow> ?test p \<Longrightarrow> ?Inv(?step p)"
-    apply(tactic {* split_all_tac @{context} 1 *})
-    using inv by iprover
-  moreover have "?Inv ([x],{x})" by (simp)
-  ultimately have I: "?Inv (ws,Z)"
+proof -
+  have "rtrancl_while_invariant ([x],{x})" by simp
+  with rtrancl_while_invariant have I: "rtrancl_while_invariant (ws,Z)"
     by (rule while_option_rule[OF _ assms[unfolded rtrancl_while_def]])
   { assume "ws = []"
     hence ?thesis using I
@@ -350,4 +357,41 @@
   ultimately show ?thesis by simp
 qed
 
+lemma rtrancl_while_finite_Some:
+  assumes "finite ({(x, y). y \<in> set (f x)}\<^sup>* `` {x})" (is "finite ?Cl")
+  shows "\<exists>y. rtrancl_while = Some y"
+proof -
+  let ?R = "(\<lambda>(_, Z). card (?Cl - Z)) <*mlex*> (\<lambda>(ws, _). length ws) <*mlex*> {}"
+  have "wf ?R" by (blast intro: wf_mlex)
+  then show ?thesis unfolding rtrancl_while_def
+  proof (rule wf_rel_while_option_Some[of ?R rtrancl_while_invariant])
+    fix st assume *: "rtrancl_while_invariant st \<and> rtrancl_while_test st"
+    hence I: "rtrancl_while_invariant (rtrancl_while_step st)"
+      by (blast intro: rtrancl_while_invariant)
+    show "(rtrancl_while_step st, st) \<in> ?R"
+    proof (cases st)
+      fix ws Z let ?ws = "fst (rtrancl_while_step st)" and ?Z = "snd (rtrancl_while_step st)"
+      assume st: "st = (ws, Z)"
+      with * obtain h t where ws: "ws = h # t" "p h" by (cases ws) auto
+      { assume "remdups (filter (\<lambda>y. y \<notin> Z) (f h)) \<noteq> []"
+        then obtain z where "z \<in> set (remdups (filter (\<lambda>y. y \<notin> Z) (f h)))" by fastforce
+        with st ws I have "Z \<subset> ?Z" "Z \<subseteq> ?Cl" "?Z \<subseteq> ?Cl" by auto
+        with assms have "card (?Cl - ?Z) < card (?Cl - Z)" by (blast intro: psubset_card_mono)
+        with st ws have ?thesis unfolding mlex_prod_def by simp
+      }
+      moreover
+      { assume "remdups (filter (\<lambda>y. y \<notin> Z) (f h)) = []"
+        with st ws have "?Z = Z" "?ws = t"  by (auto simp: filter_empty_conv)
+        with st ws have ?thesis unfolding mlex_prod_def by simp
+      }
+      ultimately show ?thesis by blast
+    qed
+  qed (simp_all add: rtrancl_while_invariant)
+qed
+
 end
+
+hide_const (open) rtrancl_while_test rtrancl_while_step rtrancl_while_invariant
+hide_fact (open) rtrancl_while_invariant
+
+end
--- a/src/HOL/Lifting_Set.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Lifting_Set.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -153,7 +153,7 @@
   unfolding fun_rel_def set_rel_def by fast
 
 lemma SUPR_parametric [transfer_rule]:
-  "(set_rel R ===> (R ===> op =) ===> op =) SUPR SUPR"
+  "(set_rel R ===> (R ===> op =) ===> op =) SUPR (SUPR :: _ \<Rightarrow> _ \<Rightarrow> _::complete_lattice)"
 proof(rule fun_relI)+
   fix A B f and g :: "'b \<Rightarrow> 'c"
   assume AB: "set_rel R A B"
--- a/src/HOL/Limits.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Limits.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -138,6 +138,18 @@
 lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X"
 by (auto simp add: Bseq_def)
 
+lemma Bseq_bdd_above: "Bseq (X::nat \<Rightarrow> real) \<Longrightarrow> bdd_above (range X)"
+proof (elim BseqE, intro bdd_aboveI2)
+  fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "X n \<le> K"
+    by (auto elim!: allE[of _ n])
+qed
+
+lemma Bseq_bdd_below: "Bseq (X::nat \<Rightarrow> real) \<Longrightarrow> bdd_below (range X)"
+proof (elim BseqE, intro bdd_belowI2)
+  fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "- K \<le> X n"
+    by (auto elim!: allE[of _ n])
+qed
+
 lemma lemma_NBseq_def:
   "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
 proof safe
@@ -179,7 +191,7 @@
 apply (rule_tac x = K in exI, simp)
 apply (rule exI [where x = 0], auto)
 apply (erule order_less_le_trans, simp)
-apply (drule_tac x=n in spec, fold diff_minus)
+apply (drule_tac x=n in spec)
 apply (drule order_trans [OF norm_triangle_ineq2])
 apply simp
 done
@@ -192,9 +204,11 @@
   then obtain K
     where *: "0 < K" and **: "\<And>n. norm (X n) \<le> K" by (auto simp add: Bseq_def)
   from * have "0 < K + norm (X 0)" by (rule order_less_le_trans) simp
-  moreover from ** have "\<forall>n. norm (X n + - X 0) \<le> K + norm (X 0)"
-    by (auto intro: order_trans norm_triangle_ineq)
-  ultimately show ?Q by blast
+  from ** have "\<forall>n. norm (X n - X 0) \<le> K + norm (X 0)"
+    by (auto intro: order_trans norm_triangle_ineq4)
+  then have "\<forall>n. norm (X n + - X 0) \<le> K + norm (X 0)"
+    by simp
+  with `0 < K + norm (X 0)` show ?Q by blast
 next
   assume ?Q then show ?P by (auto simp add: Bseq_iff2)
 qed
@@ -205,20 +219,9 @@
 apply (drule_tac x = n in spec, arith)
 done
 
+
 subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
 
-lemma Bseq_isUb:
-  "!!(X::nat=>real). Bseq X ==> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U"
-by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff)
-
-text{* Use completeness of reals (supremum property)
-   to show that any bounded sequence has a least upper bound*}
-
-lemma Bseq_isLub:
-  "!!(X::nat=>real). Bseq X ==>
-   \<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U"
-by (blast intro: reals_complete Bseq_isUb)
-
 lemma Bseq_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X"
   by (simp add: Bseq_def)
 
@@ -342,7 +345,7 @@
   unfolding Zfun_def by simp
 
 lemma Zfun_diff: "\<lbrakk>Zfun f F; Zfun g F\<rbrakk> \<Longrightarrow> Zfun (\<lambda>x. f x - g x) F"
-  by (simp only: diff_minus Zfun_add Zfun_minus)
+  using Zfun_add [of f F "\<lambda>x. - g x"] by (simp add: Zfun_minus)
 
 lemma (in bounded_linear) Zfun:
   assumes g: "Zfun g F"
@@ -532,7 +535,7 @@
 lemma tendsto_diff [tendsto_intros]:
   fixes a b :: "'a::real_normed_vector"
   shows "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x - g x) ---> a - b) F"
-  by (simp add: diff_minus tendsto_add tendsto_minus)
+  using tendsto_add [of f a F "\<lambda>x. - g x" "- b"] by (simp add: tendsto_minus)
 
 lemma continuous_diff [continuous_intros]:
   fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
@@ -1355,40 +1358,29 @@
 
 text {* A monotone sequence converges to its least upper bound. *}
 
-lemma isLub_mono_imp_LIMSEQ:
-  fixes X :: "nat \<Rightarrow> real"
-  assumes u: "isLub UNIV {x. \<exists>n. X n = x} u" (* FIXME: use 'range X' *)
-  assumes X: "\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n"
-  shows "X ----> u"
-proof (rule LIMSEQ_I)
-  have 1: "\<forall>n. X n \<le> u"
-    using isLubD2 [OF u] by auto
-  have "\<forall>y. (\<forall>n. X n \<le> y) \<longrightarrow> u \<le> y"
-    using isLub_le_isUb [OF u] by (auto simp add: isUb_def setle_def)
-  hence 2: "\<forall>y<u. \<exists>n. y < X n"
-    by (metis not_le)
-  fix r :: real assume "0 < r"
-  hence "u - r < u" by simp
-  hence "\<exists>m. u - r < X m" using 2 by simp
-  then obtain m where "u - r < X m" ..
-  with X have "\<forall>n\<ge>m. u - r < X n"
-    by (fast intro: less_le_trans)
-  hence "\<exists>m. \<forall>n\<ge>m. u - r < X n" ..
-  thus "\<exists>m. \<forall>n\<ge>m. norm (X n - u) < r"
-    using 1 by (simp add: diff_less_eq add_commute)
-qed
+lemma LIMSEQ_incseq_SUP:
+  fixes X :: "nat \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology}"
+  assumes u: "bdd_above (range X)"
+  assumes X: "incseq X"
+  shows "X ----> (SUP i. X i)"
+  by (rule order_tendstoI)
+     (auto simp: eventually_sequentially u less_cSUP_iff intro: X[THEN incseqD] less_le_trans cSUP_lessD[OF u])
 
-text{*A standard proof of the theorem for monotone increasing sequence*}
-
-lemma Bseq_mono_convergent:
-   "Bseq X \<Longrightarrow> \<forall>m. \<forall>n \<ge> m. X m \<le> X n \<Longrightarrow> convergent (X::nat=>real)"
-  by (metis Bseq_isLub isLub_mono_imp_LIMSEQ convergentI)
+lemma LIMSEQ_decseq_INF:
+  fixes X :: "nat \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology}"
+  assumes u: "bdd_below (range X)"
+  assumes X: "decseq X"
+  shows "X ----> (INF i. X i)"
+  by (rule order_tendstoI)
+     (auto simp: eventually_sequentially u cINF_less_iff intro: X[THEN decseqD] le_less_trans less_cINF_D[OF u])
 
 text{*Main monotonicity theorem*}
 
 lemma Bseq_monoseq_convergent: "Bseq X \<Longrightarrow> monoseq X \<Longrightarrow> convergent (X::nat\<Rightarrow>real)"
-  by (metis monoseq_iff incseq_def decseq_eq_incseq convergent_minus_iff Bseq_minus_iff
-            Bseq_mono_convergent)
+  by (auto simp: monoseq_iff convergent_def intro: LIMSEQ_decseq_INF LIMSEQ_incseq_SUP dest: Bseq_bdd_above Bseq_bdd_below)
+
+lemma Bseq_mono_convergent: "Bseq X \<Longrightarrow> (\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n) \<Longrightarrow> convergent (X::nat\<Rightarrow>real)"
+  by (auto intro!: Bseq_monoseq_convergent incseq_imp_monoseq simp: incseq_def)
 
 lemma Cauchy_iff:
   fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
--- a/src/HOL/List.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/List.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -902,7 +902,7 @@
 lemma self_append_conv [iff]: "(xs = xs @ ys) = (ys = [])"
 by (induct xs) auto
 
-lemma append_eq_append_conv [simp, no_atp]:
+lemma append_eq_append_conv [simp]:
  "length xs = length ys \<or> length us = length vs
  ==> (xs@us = ys@vs) = (xs=ys \<and> us=vs)"
 apply (induct xs arbitrary: ys)
@@ -934,7 +934,7 @@
 lemma self_append_conv2 [iff]: "(ys = xs @ ys) = (xs = [])"
 using append_same_eq [of "[]"] by auto
 
-lemma hd_Cons_tl [simp,no_atp]: "xs \<noteq> [] ==> hd xs # tl xs = xs"
+lemma hd_Cons_tl [simp]: "xs \<noteq> [] ==> hd xs # tl xs = xs"
 by (induct xs) auto
 
 lemma hd_append: "hd (xs @ ys) = (if xs = [] then hd ys else hd xs)"
@@ -1178,7 +1178,7 @@
 lemma singleton_rev_conv [simp]: "([x] = rev xs) = (xs = [x])"
 by (cases xs) auto
 
-lemma rev_is_rev_conv [iff, no_atp]: "(rev xs = rev ys) = (xs = ys)"
+lemma rev_is_rev_conv [iff]: "(rev xs = rev ys) = (xs = ys)"
 apply (induct xs arbitrary: ys, force)
 apply (case_tac ys, simp, force)
 done
@@ -5084,10 +5084,10 @@
   for A :: "'a set"
 where
     Nil [intro!, simp]: "[]: lists A"
-  | Cons [intro!, simp, no_atp]: "[| a: A; l: lists A|] ==> a#l : lists A"
-
-inductive_cases listsE [elim!,no_atp]: "x#l : lists A"
-inductive_cases listspE [elim!,no_atp]: "listsp A (x # l)"
+  | Cons [intro!, simp]: "[| a: A; l: lists A|] ==> a#l : lists A"
+
+inductive_cases listsE [elim!]: "x#l : lists A"
+inductive_cases listspE [elim!]: "listsp A (x # l)"
 
 inductive_simps listsp_simps[code]:
   "listsp A []"
@@ -5129,15 +5129,15 @@
 
 lemmas in_lists_conv_set [code_unfold] = in_listsp_conv_set [to_set]
 
-lemma in_listspD [dest!,no_atp]: "listsp A xs ==> \<forall>x\<in>set xs. A x"
+lemma in_listspD [dest!]: "listsp A xs ==> \<forall>x\<in>set xs. A x"
 by (rule in_listsp_conv_set [THEN iffD1])
 
-lemmas in_listsD [dest!,no_atp] = in_listspD [to_set]
-
-lemma in_listspI [intro!,no_atp]: "\<forall>x\<in>set xs. A x ==> listsp A xs"
+lemmas in_listsD [dest!] = in_listspD [to_set]
+
+lemma in_listspI [intro!]: "\<forall>x\<in>set xs. A x ==> listsp A xs"
 by (rule in_listsp_conv_set [THEN iffD2])
 
-lemmas in_listsI [intro!,no_atp] = in_listspI [to_set]
+lemmas in_listsI [intro!] = in_listspI [to_set]
 
 lemma lists_eq_set: "lists A = {xs. set xs <= A}"
 by auto
@@ -5514,15 +5514,15 @@
 text{* Accessible part and wellfoundedness: *}
 
 lemma Cons_acc_listrel1I [intro!]:
-  "x \<in> acc r \<Longrightarrow> xs \<in> acc (listrel1 r) \<Longrightarrow> (x # xs) \<in> acc (listrel1 r)"
-apply (induct arbitrary: xs set: acc)
+  "x \<in> Wellfounded.acc r \<Longrightarrow> xs \<in> Wellfounded.acc (listrel1 r) \<Longrightarrow> (x # xs) \<in> Wellfounded.acc (listrel1 r)"
+apply (induct arbitrary: xs set: Wellfounded.acc)
 apply (erule thin_rl)
 apply (erule acc_induct)
 apply (rule accI)
 apply (blast)
 done
 
-lemma lists_accD: "xs \<in> lists (acc r) \<Longrightarrow> xs \<in> acc (listrel1 r)"
+lemma lists_accD: "xs \<in> lists (Wellfounded.acc r) \<Longrightarrow> xs \<in> Wellfounded.acc (listrel1 r)"
 apply (induct set: lists)
  apply (rule accI)
  apply simp
@@ -5530,8 +5530,8 @@
 apply (fast dest: acc_downward)
 done
 
-lemma lists_accI: "xs \<in> acc (listrel1 r) \<Longrightarrow> xs \<in> lists (acc r)"
-apply (induct set: acc)
+lemma lists_accI: "xs \<in> Wellfounded.acc (listrel1 r) \<Longrightarrow> xs \<in> lists (Wellfounded.acc r)"
+apply (induct set: Wellfounded.acc)
 apply clarify
 apply (rule accI)
 apply (fastforce dest!: in_set_conv_decomp[THEN iffD1] simp: listrel1_def)
--- a/src/HOL/Lubs.thy	Mon Nov 11 17:34:44 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,103 +0,0 @@
-(*  Title:      HOL/Lubs.thy
-    Author:     Jacques D. Fleuriot, University of Cambridge
-*)
-
-header {* Definitions of Upper Bounds and Least Upper Bounds *}
-
-theory Lubs
-imports Main
-begin
-
-text {* Thanks to suggestions by James Margetson *}
-
-definition setle :: "'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"  (infixl "*<=" 70)
-  where "S *<= x = (ALL y: S. y \<le> x)"
-
-definition setge :: "'a::ord \<Rightarrow> 'a set \<Rightarrow> bool"  (infixl "<=*" 70)
-  where "x <=* S = (ALL y: S. x \<le> y)"
-
-definition leastP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a::ord \<Rightarrow> bool"
-  where "leastP P x = (P x \<and> x <=* Collect P)"
-
-definition isUb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
-  where "isUb R S x = (S *<= x \<and> x: R)"
-
-definition isLub :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
-  where "isLub R S x = leastP (isUb R S) x"
-
-definition ubs :: "'a set \<Rightarrow> 'a::ord set \<Rightarrow> 'a set"
-  where "ubs R S = Collect (isUb R S)"
-
-
-subsection {* Rules for the Relations @{text "*<="} and @{text "<=*"} *}
-
-lemma setleI: "ALL y: S. y \<le> x \<Longrightarrow> S *<= x"
-  by (simp add: setle_def)
-
-lemma setleD: "S *<= x \<Longrightarrow> y: S \<Longrightarrow> y \<le> x"
-  by (simp add: setle_def)
-
-lemma setgeI: "ALL y: S. x \<le> y \<Longrightarrow> x <=* S"
-  by (simp add: setge_def)
-
-lemma setgeD: "x <=* S \<Longrightarrow> y: S \<Longrightarrow> x \<le> y"
-  by (simp add: setge_def)
-
-
-subsection {* Rules about the Operators @{term leastP}, @{term ub} and @{term lub} *}
-
-lemma leastPD1: "leastP P x \<Longrightarrow> P x"
-  by (simp add: leastP_def)
-
-lemma leastPD2: "leastP P x \<Longrightarrow> x <=* Collect P"
-  by (simp add: leastP_def)
-
-lemma leastPD3: "leastP P x \<Longrightarrow> y: Collect P \<Longrightarrow> x \<le> y"
-  by (blast dest!: leastPD2 setgeD)
-
-lemma isLubD1: "isLub R S x \<Longrightarrow> S *<= x"
-  by (simp add: isLub_def isUb_def leastP_def)
-
-lemma isLubD1a: "isLub R S x \<Longrightarrow> x: R"
-  by (simp add: isLub_def isUb_def leastP_def)
-
-lemma isLub_isUb: "isLub R S x \<Longrightarrow> isUb R S x"
-  unfolding isUb_def by (blast dest: isLubD1 isLubD1a)
-
-lemma isLubD2: "isLub R S x \<Longrightarrow> y : S \<Longrightarrow> y \<le> x"
-  by (blast dest!: isLubD1 setleD)
-
-lemma isLubD3: "isLub R S x \<Longrightarrow> leastP (isUb R S) x"
-  by (simp add: isLub_def)
-
-lemma isLubI1: "leastP(isUb R S) x \<Longrightarrow> isLub R S x"
-  by (simp add: isLub_def)
-
-lemma isLubI2: "isUb R S x \<Longrightarrow> x <=* Collect (isUb R S) \<Longrightarrow> isLub R S x"
-  by (simp add: isLub_def leastP_def)
-
-lemma isUbD: "isUb R S x \<Longrightarrow> y : S \<Longrightarrow> y \<le> x"
-  by (simp add: isUb_def setle_def)
-
-lemma isUbD2: "isUb R S x \<Longrightarrow> S *<= x"
-  by (simp add: isUb_def)
-
-lemma isUbD2a: "isUb R S x \<Longrightarrow> x: R"
-  by (simp add: isUb_def)
-
-lemma isUbI: "S *<= x \<Longrightarrow> x: R \<Longrightarrow> isUb R S x"
-  by (simp add: isUb_def)
-
-lemma isLub_le_isUb: "isLub R S x \<Longrightarrow> isUb R S y \<Longrightarrow> x \<le> y"
-  unfolding isLub_def by (blast intro!: leastPD3)
-
-lemma isLub_ubs: "isLub R S x \<Longrightarrow> x <=* ubs R S"
-  unfolding ubs_def isLub_def by (rule leastPD2)
-
-lemma isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::'a::linorder)"
-  apply (frule isLub_isUb)
-  apply (frule_tac x = y in isLub_isUb)
-  apply (blast intro!: order_antisym dest!: isLub_le_isUb)
-  done
-
-end
--- a/src/HOL/Matrix_LP/LP.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Matrix_LP/LP.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -72,8 +72,7 @@
 proof -
   have "0 <= A - A1"    
   proof -
-    have 1: "A - A1 = A + (- A1)" by simp
-    show ?thesis by (simp only: 1 add_right_mono[of A1 A "-A1", simplified, simplified assms])
+    from assms add_right_mono [of A1 A "- A1"] show ?thesis by simp
   qed
   then have "abs (A-A1) = A-A1" by (rule abs_of_nonneg)
   with assms show "abs (A-A1) <= (A2-A1)" by simp
@@ -147,9 +146,9 @@
   then have "c * x <= y * b - (y * A - c) * x" by (simp add: le_diff_eq)
   then have cx: "c * x <= y * b + (c - y * A) * x" by (simp add: algebra_simps)
   have s2: "c - y * A <= c2 - y * A1"
-    by (simp add: diff_minus assms add_mono mult_left_mono)
+    by (simp add: assms add_mono mult_left_mono algebra_simps)
   have s1: "c1 - y * A2 <= c - y * A"
-    by (simp add: diff_minus assms add_mono mult_left_mono)
+    by (simp add: assms add_mono mult_left_mono algebra_simps)
   have prts: "(c - y * A) * x <= ?C"
     apply (simp add: Let_def)
     apply (rule mult_le_prts)
--- a/src/HOL/Matrix_LP/Matrix.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Matrix_LP/Matrix.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1542,8 +1542,8 @@
   fix A B :: "'a matrix"
   show "- A + A = 0" 
     by (simp add: plus_matrix_def minus_matrix_def Rep_matrix_inject[symmetric] ext)
-  show "A - B = A + - B"
-    by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def Rep_matrix_inject [symmetric] ext diff_minus)
+  show "A + - B = A - B"
+    by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def Rep_matrix_inject [symmetric] ext)
 qed
 
 instance matrix :: (ab_group_add) ab_group_add
--- a/src/HOL/Meson.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Meson.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -132,45 +132,45 @@
 text{* Combinator translation helpers *}
 
 definition COMBI :: "'a \<Rightarrow> 'a" where
-[no_atp]: "COMBI P = P"
+"COMBI P = P"
 
 definition COMBK :: "'a \<Rightarrow> 'b \<Rightarrow> 'a" where
-[no_atp]: "COMBK P Q = P"
+"COMBK P Q = P"
 
-definition COMBB :: "('b => 'c) \<Rightarrow> ('a => 'b) \<Rightarrow> 'a \<Rightarrow> 'c" where [no_atp]:
+definition COMBB :: "('b => 'c) \<Rightarrow> ('a => 'b) \<Rightarrow> 'a \<Rightarrow> 'c" where
 "COMBB P Q R = P (Q R)"
 
 definition COMBC :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'c" where
-[no_atp]: "COMBC P Q R = P R Q"
+"COMBC P Q R = P R Q"
 
 definition COMBS :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" where
-[no_atp]: "COMBS P Q R = P R (Q R)"
+"COMBS P Q R = P R (Q R)"
 
-lemma abs_S [no_atp]: "\<lambda>x. (f x) (g x) \<equiv> COMBS f g"
+lemma abs_S: "\<lambda>x. (f x) (g x) \<equiv> COMBS f g"
 apply (rule eq_reflection)
 apply (rule ext) 
 apply (simp add: COMBS_def) 
 done
 
-lemma abs_I [no_atp]: "\<lambda>x. x \<equiv> COMBI"
+lemma abs_I: "\<lambda>x. x \<equiv> COMBI"
 apply (rule eq_reflection)
 apply (rule ext) 
 apply (simp add: COMBI_def) 
 done
 
-lemma abs_K [no_atp]: "\<lambda>x. y \<equiv> COMBK y"
+lemma abs_K: "\<lambda>x. y \<equiv> COMBK y"
 apply (rule eq_reflection)
 apply (rule ext) 
 apply (simp add: COMBK_def) 
 done
 
-lemma abs_B [no_atp]: "\<lambda>x. a (g x) \<equiv> COMBB a g"
+lemma abs_B: "\<lambda>x. a (g x) \<equiv> COMBB a g"
 apply (rule eq_reflection)
 apply (rule ext) 
 apply (simp add: COMBB_def) 
 done
 
-lemma abs_C [no_atp]: "\<lambda>x. (f x) b \<equiv> COMBC f b"
+lemma abs_C: "\<lambda>x. (f x) b \<equiv> COMBC f b"
 apply (rule eq_reflection)
 apply (rule ext) 
 apply (simp add: COMBC_def) 
@@ -180,7 +180,7 @@
 subsection {* Skolemization helpers *}
 
 definition skolem :: "'a \<Rightarrow> 'a" where
-[no_atp]: "skolem = (\<lambda>x. x)"
+"skolem = (\<lambda>x. x)"
 
 lemma skolem_COMBK_iff: "P \<longleftrightarrow> skolem (COMBK P (i\<Colon>nat))"
 unfolding skolem_def COMBK_def by (rule refl)
--- a/src/HOL/Metis.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Metis.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -16,7 +16,7 @@
 subsection {* Literal selection and lambda-lifting helpers *}
 
 definition select :: "'a \<Rightarrow> 'a" where
-[no_atp]: "select = (\<lambda>x. x)"
+"select = (\<lambda>x. x)"
 
 lemma not_atomize: "(\<not> A \<Longrightarrow> False) \<equiv> Trueprop A"
 by (cut_tac atomize_not [of "\<not> A"]) simp
@@ -30,7 +30,7 @@
 lemma select_FalseI: "False \<Longrightarrow> select False" by simp
 
 definition lambda :: "'a \<Rightarrow> 'a" where
-[no_atp]: "lambda = (\<lambda>x. x)"
+"lambda = (\<lambda>x. x)"
 
 lemma eq_lambdaI: "x \<equiv> y \<Longrightarrow> x \<equiv> lambda y"
 unfolding lambda_def by assumption
--- a/src/HOL/Metis_Examples/Big_O.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Metis_Examples/Big_O.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -493,8 +493,10 @@
 
 lemma bigo_compose2:
 "f =o g +o O(h) \<Longrightarrow> (\<lambda>x. f(k x)) =o (\<lambda>x. g(k x)) +o O(\<lambda>x. h(k x))"
-apply (simp only: set_minus_plus [symmetric] diff_minus fun_Compl_def func_plus)
-by (erule bigo_compose1)
+apply (simp only: set_minus_plus [symmetric] fun_Compl_def func_plus)
+apply (drule bigo_compose1 [of "f - g" h k])
+apply (simp add: fun_diff_def)
+done
 
 subsection {* Setsum *}
 
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -439,7 +439,7 @@
                   term_order |> the_default I)
             #> (Option.map (Config.put ATP_Systems.force_sos)
                   force_sos |> the_default I))
-    val params as {max_facts, slice, ...} =
+    val params as {max_facts, ...} =
       Sledgehammer_Isar.default_params ctxt
          ([("verbose", "true"),
            ("fact_filter", fact_filter),
@@ -454,8 +454,7 @@
           |> sh_minimizeLST (*don't confuse the two minimization flags*)
           |> max_new_mono_instancesLST
           |> max_mono_itersLST)
-    val default_max_facts =
-      Sledgehammer_Provers.default_max_facts_of_prover ctxt slice prover_name
+    val default_max_facts = Sledgehammer_Provers.default_max_facts_of_prover ctxt prover_name
     val is_appropriate_prop =
       Sledgehammer_Provers.is_appropriate_prop_of_prover ctxt prover_name
     val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal goal i ctxt
@@ -494,7 +493,7 @@
                      |> Output.urgent_message)
         val prover = get_prover ctxt prover_name params goal facts
         val problem =
-          {state = st', goal = goal, subgoal = i,
+          {comment = "", state = st', goal = goal, subgoal = i,
            subgoal_count = Sledgehammer_Util.subgoal_count st, factss = factss}
       in prover params (K (K (K ""))) problem end)) ()
       handle TimeLimit.TimeOut => failed ATP_Proof.TimedOut
@@ -584,6 +583,7 @@
         ({pre=st, log, ...}: Mirabelle.run_args) =
   let
     val ctxt = Proof.context_of st
+    val {goal, ...} = Proof.goal st
     val n0 = length (these (!named_thms))
     val prover_name = get_prover_name ctxt args
     val type_enc = AList.lookup (op =) args type_encK |> the_default type_enc_default
@@ -613,7 +613,7 @@
           true 1 (Sledgehammer_Util.subgoal_count st)
     val _ = log separator
     val (used_facts, (preplay, message, message_tail)) =
-      minimize st NONE (these (!named_thms))
+      minimize st goal NONE (these (!named_thms))
     val msg = message (Lazy.force preplay) ^ message_tail
   in
     case used_facts of
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -110,18 +110,14 @@
        SOME proofs =>
        let
          val {context = ctxt, facts = chained_ths, goal} = Proof.goal pre
-         val prover = AList.lookup (op =) args "prover"
-                      |> the_default default_prover
-         val params as {max_facts, slice, ...} =
+         val prover = AList.lookup (op =) args "prover" |> the_default default_prover
+         val params as {max_facts, ...} =
            Sledgehammer_Isar.default_params ctxt args
-         val default_max_facts =
-           Sledgehammer_Provers.default_max_facts_of_prover ctxt slice prover
+         val default_max_facts = Sledgehammer_Provers.default_max_facts_of_prover ctxt prover
          val is_appropriate_prop =
-           Sledgehammer_Provers.is_appropriate_prop_of_prover ctxt
-               default_prover
+           Sledgehammer_Provers.is_appropriate_prop_of_prover ctxt default_prover
          val relevance_fudge =
-           extract_relevance_fudge args
-               (Sledgehammer_Provers.relevance_fudge_of_prover ctxt prover)
+           extract_relevance_fudge args Sledgehammer_MePo.default_relevance_fudge
          val subgoal = 1
          val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal goal subgoal ctxt
          val ho_atp = Sledgehammer_Provers.is_ho_atp ctxt prover
@@ -132,9 +128,9 @@
                Sledgehammer_Fact.no_fact_override reserved css_table chained_ths
                hyp_ts concl_t
            |> filter (is_appropriate_prop o prop_of o snd)
+           |> Sledgehammer_Fact.drop_duplicate_facts
            |> Sledgehammer_MePo.mepo_suggested_facts ctxt params
-                  default_prover (the_default default_max_facts max_facts)
-                  (SOME relevance_fudge) hyp_ts concl_t
+                  (the_default default_max_facts max_facts) (SOME relevance_fudge) hyp_ts concl_t
             |> map (fst o fst)
          val (found_facts, lost_facts) =
            flat proofs |> sort_distinct string_ord
--- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1099,8 +1099,8 @@
   shows "(\<lambda>x. m *s x + c) o (\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) = id"
   "(\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) o (\<lambda>x. m *s x + c) = id"
   using m0
-  apply (auto simp add: fun_eq_iff vector_add_ldistrib)
-  apply (simp_all add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1[symmetric])
+  apply (auto simp add: fun_eq_iff vector_add_ldistrib diff_conv_add_uminus simp del: add_uminus_conv_diff)
+  apply (simp_all add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1 [symmetric])
   done
 
 lemma vector_affinity_eq:
@@ -1114,7 +1114,7 @@
     using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
 next
   assume h: "x = inverse m *s y + - (inverse m *s c)"
-  show "m *s x + c = y" unfolding h diff_minus[symmetric]
+  show "m *s x + c = y" unfolding h
     using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
 qed
 
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -858,9 +858,10 @@
   assumes "affine_parallel A B"
   shows "affine_parallel B A"
 proof -
-  from assms obtain a where "B = (\<lambda>x. a + x) ` A"
+  from assms obtain a where B: "B = (\<lambda>x. a + x) ` A"
     unfolding affine_parallel_def by auto
-  then show ?thesis
+  have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff)
+  from B show ?thesis
     using translation_galois [of B a A]
     unfolding affine_parallel_def by auto
 qed
@@ -980,6 +981,7 @@
   assumes "affine S" "a \<in> S"
   shows "subspace ((\<lambda>x. (-a)+x) ` S)"
 proof -
+  have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff)
   have "affine ((\<lambda>x. (-a)+x) ` S)"
     using  affine_translation assms by auto
   moreover have "0 : ((\<lambda>x. (-a)+x) ` S)"
@@ -992,15 +994,12 @@
   assumes "L \<equiv> {y. \<exists>x \<in> S. (-a)+x=y}"
   shows "subspace L & affine_parallel S L"
 proof -
-  have par: "affine_parallel S L"
-    unfolding affine_parallel_def using assms by auto
+  from assms have "L = plus (- a) ` S" by auto
+  then have par: "affine_parallel S L"
+    unfolding affine_parallel_def .. 
   then have "affine L" using assms parallel_is_affine by auto
   moreover have "0 \<in> L"
-    using assms
-    apply auto
-    using exI[of "(\<lambda>x. x:S \<and> -a+x=0)" a]
-    apply auto
-    done
+    using assms by auto
   ultimately show ?thesis
     using subspace_affine par by auto
 qed
@@ -2390,7 +2389,7 @@
   ultimately have h1: "affine hull ((\<lambda>x. a + x) `  S) \<subseteq> (\<lambda>x. a + x) ` (affine hull S)"
     by (metis hull_minimal)
   have "affine((\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) `  S)))"
-    using affine_translation affine_affine_hull by auto
+    using affine_translation affine_affine_hull by (auto simp del: uminus_add_conv_diff)
   moreover have "(\<lambda>x. -a + x) ` (\<lambda>x. a + x) `  S \<subseteq> (\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) `  S))"
     using hull_subset[of "(\<lambda>x. a + x) `  S"] by auto
   moreover have "S = (\<lambda>x. -a + x) ` (\<lambda>x. a + x) `  S"
@@ -2478,7 +2477,7 @@
     using affine_dependent_translation_eq[of "(insert a S)" "-a"]
       affine_dependent_imp_dependent2 assms
       dependent_imp_affine_dependent[of a S]
-    by auto
+    by (auto simp del: uminus_add_conv_diff)
 qed
 
 lemma affine_dependent_iff_dependent2:
@@ -2512,7 +2511,7 @@
     then have "insert 0 ((\<lambda>x. -a+x) ` (s - {a})) = (\<lambda>x. -a+x) ` s"
       by auto
     then have "span ((\<lambda>x. -a+x) ` (s - {a}))=span ((\<lambda>x. -a+x) ` s)"
-      using span_insert_0[of "op + (- a) ` (s - {a})"] by auto
+      using span_insert_0[of "op + (- a) ` (s - {a})"] by (auto simp del: uminus_add_conv_diff)
     moreover have "{x - a |x. x \<in> (s - {a})} = ((\<lambda>x. -a+x) ` (s - {a}))"
       by auto
     moreover have "insert a (s - {a}) = insert a s"
@@ -2652,7 +2651,7 @@
     moreover have h1: "card ((\<lambda>x. -a + x) ` (B-{a})) = card (B-{a})"
        apply (rule card_image)
        using translate_inj_on
-       apply auto
+       apply (auto simp del: uminus_add_conv_diff)
        done
     ultimately have "card (B-{a}) > 0" by auto
     then have *: "finite (B - {a})"
@@ -4507,38 +4506,30 @@
     apply (erule_tac x="x - y" in ballE)
     apply (auto simp add: inner_diff)
     done
-  def k \<equiv> "Sup ((\<lambda>x. inner a x) ` t)"
+  def k \<equiv> "SUP x:t. a \<bullet> x"
   show ?thesis
     apply (rule_tac x="-a" in exI)
     apply (rule_tac x="-(k + b / 2)" in exI)
-    apply rule
-    apply rule
-    defer
-    apply rule
+    apply (intro conjI ballI)
     unfolding inner_minus_left and neg_less_iff_less
   proof -
-    from ab have "((\<lambda>x. inner a x) ` t) *<= (inner a y - b)"
-      apply (erule_tac x=y in ballE)
-      apply (rule setleI)
-      using `y \<in> s`
-      apply auto
-      done
-    then have k: "isLub UNIV ((\<lambda>x. inner a x) ` t) k"
+    fix x assume "x \<in> t"
+    then have "inner a x - b / 2 < k"
       unfolding k_def
-      apply (rule_tac isLub_cSup)
-      using assms(5)
-      apply auto
-      done
-    fix x
-    assume "x \<in> t"
-    then show "inner a x < (k + b / 2)"
-      using `0<b` and isLubD2[OF k, of "inner a x"] by auto
+    proof (subst less_cSUP_iff)
+      show "t \<noteq> {}" by fact
+      show "bdd_above (op \<bullet> a ` t)"
+        using ab[rule_format, of y] `y \<in> s`
+        by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le)
+    qed (auto intro!: bexI[of _ x] `0<b`)
+    then show "inner a x < k + b / 2"
+      by auto
   next
     fix x
     assume "x \<in> s"
     then have "k \<le> inner a x - b"
       unfolding k_def
-      apply (rule_tac cSup_least)
+      apply (rule_tac cSUP_least)
       using assms(5)
       using ab[THEN bspec[where x=x]]
       apply auto
@@ -4627,20 +4618,14 @@
   from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]]
   obtain a where "a \<noteq> 0" "\<forall>x\<in>{x - y |x y. x \<in> t \<and> y \<in> s}. 0 \<le> inner a x"
     using assms(3-5) by auto
-  then have "\<forall>x\<in>t. \<forall>y\<in>s. inner a y \<le> inner a x"
+  then have *: "\<And>x y. x \<in> t \<Longrightarrow> y \<in> s \<Longrightarrow> inner a y \<le> inner a x"
     by (force simp add: inner_diff)
-  then show ?thesis
-    apply (rule_tac x=a in exI)
-    apply (rule_tac x="Sup ((\<lambda>x. inner a x) ` s)" in exI)
+  then have bdd: "bdd_above ((op \<bullet> a)`s)"
+    using `t \<noteq> {}` by (auto intro: bdd_aboveI2[OF *])
+  show ?thesis
     using `a\<noteq>0`
-    apply auto
-    apply (rule isLub_cSup[THEN isLubD2])
-    prefer 4
-    apply (rule cSup_least)
-    using assms(3-5)
-    apply (auto simp add: setle_def)
-    apply metis
-    done
+    by (intro exI[of _ a] exI[of _ "SUP x:s. a \<bullet> x"])
+       (auto intro!: cSUP_upper bdd cSUP_least `a \<noteq> 0` `s \<noteq> {}` *)
 qed
 
 
@@ -8725,7 +8710,7 @@
     using interior_subset[of I] `x \<in> interior I` by auto
 
   have "Inf (?F x) \<le> (f x - f y) / (x - y)"
-  proof (rule cInf_lower2)
+  proof (intro bdd_belowI cInf_lower2)
     show "(f x - f t) / (x - t) \<in> ?F x"
       using `t \<in> I` `x < t` by auto
     show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)"
--- a/src/HOL/Multivariate_Analysis/Derivative.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -516,7 +516,7 @@
       unfolding e_def
       using c[THEN conjunct1]
       using norm_minus_cancel[of "f' i - f'' i"]
-      by (auto simp add: add.commute ab_diff_minus)
+      by auto
     finally show False
       using c
       using d[THEN conjunct2,rule_format,of "x + c *\<^sub>R i"]
--- a/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -660,7 +660,7 @@
     assume "S \<noteq> {}"
     { assume ex: "\<exists>B. \<forall>x\<in>S. B \<le> x"
       then have *: "\<forall>x\<in>S. Inf S \<le> x"
-        using cInf_lower_EX[of _ S] ex by metis
+        using cInf_lower[of _ S] ex by (metis bdd_below_def)
       then have "Inf S \<in> S"
         apply (subst closed_contains_Inf)
         using ex `S \<noteq> {}` `closed S`
@@ -1193,12 +1193,12 @@
 qed
 
 lemma Liminf_at:
-  fixes f :: "'a::metric_space \<Rightarrow> _"
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice"
   shows "Liminf (at x) f = (SUP e:{0<..}. INF y:(ball x e - {x}). f y)"
   using Liminf_within[of x UNIV f] by simp
 
 lemma Limsup_at:
-  fixes f :: "'a::metric_space \<Rightarrow> _"
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice"
   shows "Limsup (at x) f = (INF e:{0<..}. SUP y:(ball x e - {x}). f y)"
   using Limsup_within[of x UNIV f] by simp
 
@@ -1209,7 +1209,7 @@
   apply (subst inf_commute)
   apply (subst SUP_inf)
   apply (intro SUP_cong[OF refl])
-  apply (cut_tac A="ball x b - {x}" and B="{x}" and M=f in INF_union)
+  apply (cut_tac A="ball x xa - {x}" and B="{x}" and M=f in INF_union)
   apply (simp add: INF_def del: inf_ereal_def)
   done
 
--- a/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -115,7 +115,7 @@
 instance vec :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add ..
 
 instance vec :: (group_add, finite) group_add
-  by default (simp_all add: vec_eq_iff diff_minus)
+  by default (simp_all add: vec_eq_iff)
 
 instance vec :: (ab_group_add, finite) ab_group_add
   by default (simp_all add: vec_eq_iff)
--- a/src/HOL/Multivariate_Analysis/Integration.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Integration.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -13,7 +13,7 @@
 lemma cSup_abs_le: (* TODO: is this really needed? *)
   fixes S :: "real set"
   shows "S \<noteq> {} \<Longrightarrow> (\<forall>x\<in>S. \<bar>x\<bar> \<le> a) \<Longrightarrow> \<bar>Sup S\<bar> \<le> a"
-  by (auto simp add: abs_le_interval_iff intro: cSup_least) (metis cSup_upper2)
+  by (auto simp add: abs_le_interval_iff intro: cSup_least) (metis cSup_upper2 bdd_aboveI)
 
 lemma cInf_abs_ge: (* TODO: is this really needed? *)
   fixes S :: "real set"
@@ -28,10 +28,10 @@
 proof -
   have th: "\<And>(x::real) l e. \<bar>x - l\<bar> \<le> e \<longleftrightarrow> l - e \<le> x \<and> x \<le> l + e"
     by arith
-  then show ?thesis
-    using S b cSup_bounds[of S "l - e" "l+e"]
-    unfolding th
-    by (auto simp add: setge_def setle_def)
+  have "bdd_above S"
+    using b by (auto intro!: bdd_aboveI[of _ "l + e"])
+  with S b show ?thesis
+    unfolding th by (auto intro!: cSup_upper2 cSup_least)
 qed
 
 lemma cInf_asclose: (* TODO: is this really needed? *)
@@ -44,9 +44,7 @@
     by auto
   also have "\<dots> \<le> e"
     apply (rule cSup_asclose)
-    apply (auto simp add: S)
-    apply (metis abs_minus_add_cancel b add_commute diff_minus)
-    done
+    using abs_minus_add_cancel b by (auto simp add: S)
   finally have "\<bar>- Sup (uminus ` S) - l\<bar> \<le> e" .
   then show ?thesis
     by (simp add: Inf_real_def)
@@ -72,39 +70,6 @@
   shows "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> a \<ge> Inf S \<longleftrightarrow> (\<exists>x\<in>S. a \<ge> x)"
   by (metis cInf_eq_Min Min_le_iff)
 
-lemma Inf: (* rename *)
-  fixes S :: "real set"
-  shows "S \<noteq> {} \<Longrightarrow> (\<exists>b. b <=* S) \<Longrightarrow> isGlb UNIV S (Inf S)"
-  by (auto simp add: isLb_def setle_def setge_def isGlb_def greatestP_def
-    intro: cInf_lower cInf_greatest)
-
-lemma real_le_inf_subset:
-  assumes "t \<noteq> {}"
-    and "t \<subseteq> s"
-    and "\<exists>b. b <=* s"
-  shows "Inf s \<le> Inf (t::real set)"
-  apply (rule isGlb_le_isLb)
-  apply (rule Inf[OF assms(1)])
-  apply (insert assms)
-  apply (erule exE)
-  apply (rule_tac x = b in exI)
-  apply (auto simp: isLb_def setge_def intro: cInf_lower cInf_greatest)
-  done
-
-lemma real_ge_sup_subset:
-  fixes t :: "real set"
-  assumes "t \<noteq> {}"
-    and "t \<subseteq> s"
-    and "\<exists>b. s *<= b"
-  shows "Sup s \<ge> Sup t"
-  apply (rule isLub_le_isUb)
-  apply (rule isLub_cSup[OF assms(1)])
-  apply (insert assms)
-  apply (erule exE)
-  apply (rule_tac x = b in exI)
-  apply (auto simp: isUb_def setle_def intro: cSup_upper cSup_least)
-  done
-
 (*declare not_less[simp] not_le[simp]*)
 
 lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib
@@ -380,7 +345,7 @@
                 using Basis_le_norm[OF k, of "?z - y"] unfolding dist_norm by auto
               then have "y\<bullet>k < a\<bullet>k"
                 using e[THEN conjunct1] k
-                by (auto simp add: field_simps as inner_Basis inner_simps)
+                by (auto simp add: field_simps abs_less_iff as inner_Basis inner_simps)
               then have "y \<notin> i"
                 unfolding ab mem_interval by (auto intro!: bexI[OF _ k])
               then show False using yi by auto
@@ -488,24 +453,24 @@
 subsection {* Bounds on intervals where they exist. *}
 
 definition interval_upperbound :: "('a::ordered_euclidean_space) set \<Rightarrow> 'a"
-  where "interval_upperbound s = (\<Sum>i\<in>Basis. Sup {a. \<exists>x\<in>s. x\<bullet>i = a} *\<^sub>R i)"
+  where "interval_upperbound s = (\<Sum>i\<in>Basis. (SUP x:s. x\<bullet>i) *\<^sub>R i)"
 
 definition interval_lowerbound :: "('a::ordered_euclidean_space) set \<Rightarrow> 'a"
-  where "interval_lowerbound s = (\<Sum>i\<in>Basis. Inf {a. \<exists>x\<in>s. x\<bullet>i = a} *\<^sub>R i)"
+  where "interval_lowerbound s = (\<Sum>i\<in>Basis. (INF x:s. x\<bullet>i) *\<^sub>R i)"
 
 lemma interval_upperbound[simp]:
   "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
     interval_upperbound {a..b} = (b::'a::ordered_euclidean_space)"
   unfolding interval_upperbound_def euclidean_representation_setsum
-  by (auto simp del: ex_simps simp add: Bex_def ex_simps[symmetric] eucl_le[where 'a='a] setle_def
-      intro!: cSup_unique)
+  by (auto simp del: ex_simps simp add: Bex_def ex_simps[symmetric] eucl_le[where 'a='a] SUP_def
+           intro!: cSup_eq)
 
 lemma interval_lowerbound[simp]:
   "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
     interval_lowerbound {a..b} = (a::'a::ordered_euclidean_space)"
   unfolding interval_lowerbound_def euclidean_representation_setsum
-  by (auto simp del: ex_simps simp add: Bex_def ex_simps[symmetric] eucl_le[where 'a='a] setge_def
-      intro!: cInf_unique)
+  by (auto simp del: ex_simps simp add: Bex_def ex_simps[symmetric] eucl_le[where 'a='a] INF_def
+           intro!: cInf_eq)
 
 lemmas interval_bounds = interval_upperbound interval_lowerbound
 
@@ -6629,7 +6594,7 @@
 lemma interval_bound_sing[simp]:
   "interval_upperbound {a} = a"
   "interval_lowerbound {a} = a"
-  unfolding interval_upperbound_def interval_lowerbound_def
+  unfolding interval_upperbound_def interval_lowerbound_def SUP_def INF_def
   by (auto simp: euclidean_representation)
 
 lemma additive_tagged_division_1:
@@ -11238,37 +11203,26 @@
 lemma bounded_variation_absolutely_integrable_interval:
   fixes f :: "'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
   assumes "f integrable_on {a..b}"
-    and "\<forall>d. d division_of {a..b} \<longrightarrow> setsum (\<lambda>k. norm(integral k f)) d \<le> B"
+    and *: "\<forall>d. d division_of {a..b} \<longrightarrow> setsum (\<lambda>k. norm(integral k f)) d \<le> B"
   shows "f absolutely_integrable_on {a..b}"
 proof -
-  let ?S = "(\<lambda>d. setsum (\<lambda>k. norm(integral k f)) d) ` {d. d division_of {a..b} }"
-  def i \<equiv> "Sup ?S"
-  have i: "isLub UNIV ?S i"
-    unfolding i_def
-    apply (rule isLub_cSup)
-    apply (rule elementary_interval)
-    defer
-    apply (rule_tac x=B in exI)
-    apply (rule setleI)
-    using assms(2)
-    apply auto
-    done
+  let ?f = "\<lambda>d. \<Sum>k\<in>d. norm (integral k f)" and ?D = "{d. d division_of {a..b}}"
+  have D_1: "?D \<noteq> {}"
+    by (rule elementary_interval[of a b]) auto
+  have D_2: "bdd_above (?f`?D)"
+    by (metis * mem_Collect_eq bdd_aboveI2)
+  note D = D_1 D_2
+  let ?S = "SUP x:?D. ?f x"
   show ?thesis
     apply rule
     apply (rule assms)
     apply rule
-    apply (subst has_integral[of _ i])
+    apply (subst has_integral[of _ ?S])
   proof safe
     case goal1
-    then have "i - e / 2 \<notin> Collect (isUb UNIV (setsum (\<lambda>k. norm (integral k f)) `
-      {d. d division_of {a..b}}))"
-      using isLub_ubs[OF i,rule_format]
-      unfolding setge_def ubs_def
-      by auto
-    then have "\<exists>y. y division_of {a..b} \<and> i - e / 2 < (\<Sum>k\<in>y. norm (integral k f))"
-      unfolding mem_Collect_eq isUb_def setle_def
-      by (simp add: not_le)
-    then guess d .. note d=conjunctD2[OF this]
+    then have "?S - e / 2 < ?S" by simp
+    then obtain d where d: "d division_of {a..b}" "?S - e / 2 < (\<Sum>k\<in>d. norm (integral k f))"
+      unfolding less_cSUP_iff[OF D] by auto
     note d' = division_ofD[OF this(1)]
 
     have "\<forall>x. \<exists>e>0. \<forall>i\<in>d. x \<notin> i \<longrightarrow> ball x e \<inter> i = {}"
@@ -11453,21 +11407,17 @@
         done
       note snd_p = division_ofD[OF division_of_tagged_division[OF p(1)]]
 
-      have *: "\<And>sni sni' sf sf'. abs (sf' - sni') < e / 2 \<longrightarrow> i - e / 2 < sni \<and> sni' \<le> i \<and>
-        sni \<le> sni' \<and> sf' = sf \<longrightarrow> abs (sf - i) < e"
+      have *: "\<And>sni sni' sf sf'. abs (sf' - sni') < e / 2 \<longrightarrow> ?S - e / 2 < sni \<and> sni' \<le> ?S \<and>
+        sni \<le> sni' \<and> sf' = sf \<longrightarrow> abs (sf - ?S) < e"
         by arith
-      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) - i) < e"
+      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) - ?S) < e"
         unfolding real_norm_def
         apply (rule *[rule_format,OF **])
         apply safe
         apply(rule d(2))
       proof -
-        case goal1
-        show ?case unfolding sum_p'
-          apply (rule isLubD2[OF i])
-          using division_of_tagged_division[OF p'']
-          apply auto
-          done
+        case goal1 show ?case
+          by (auto simp: sum_p' division_of_tagged_division[OF p''] D intro!: cSUP_upper)
       next
         case goal2
         have *: "{k \<inter> l | k l. k \<in> d \<and> l \<in> snd ` p} =
@@ -11758,18 +11708,13 @@
     and "\<forall>d. d division_of (\<Union>d) \<longrightarrow> setsum (\<lambda>k. norm (integral k f)) d \<le> B"
   shows "f absolutely_integrable_on UNIV"
 proof (rule absolutely_integrable_onI, fact, rule)
-  let ?S = "(\<lambda>d. setsum (\<lambda>k. norm(integral k f)) d) ` {d. d division_of  (\<Union>d)}"
-  def i \<equiv> "Sup ?S"
-  have i: "isLub UNIV ?S i"
-    unfolding i_def
-    apply (rule isLub_cSup)
-    apply (rule elementary_interval)
-    defer
-    apply (rule_tac x=B in exI)
-    apply (rule setleI)
-    using assms(2)
-    apply auto
-    done
+  let ?f = "\<lambda>d. \<Sum>k\<in>d. norm (integral k f)" and ?D = "{d. d division_of  (\<Union>d)}"
+  have D_1: "?D \<noteq> {}"
+    by (rule elementary_interval) auto
+  have D_2: "bdd_above (?f`?D)"
+    by (intro bdd_aboveI2[where M=B] assms(2)[rule_format]) simp
+  note D = D_1 D_2
+  let ?S = "SUP d:?D. ?f d"
   have f_int: "\<And>a b. f absolutely_integrable_on {a..b}"
     apply (rule bounded_variation_absolutely_integrable_interval[where B=B])
     apply (rule integrable_on_subinterval[OF assms(1)])
@@ -11778,7 +11723,7 @@
     apply (rule assms(2)[rule_format])
     apply auto
     done
-  show "((\<lambda>x. norm (f x)) has_integral i) UNIV"
+  show "((\<lambda>x. norm (f x)) has_integral ?S) UNIV"
     apply (subst has_integral_alt')
     apply safe
   proof -
@@ -11787,16 +11732,11 @@
       using f_int[of a b] by auto
   next
     case goal2
-    have "\<exists>y\<in>setsum (\<lambda>k. norm (integral k f)) ` {d. d division_of \<Union>d}. \<not> y \<le> i - e"
+    have "\<exists>y\<in>setsum (\<lambda>k. norm (integral k f)) ` {d. d division_of \<Union>d}. \<not> y \<le> ?S - e"
     proof (rule ccontr)
       assume "\<not> ?thesis"
-      then have "i \<le> i - e"
-        apply -
-        apply (rule isLub_le_isUb[OF i])
-        apply (rule isUbI)
-        unfolding setle_def
-        apply auto
-        done
+      then have "?S \<le> ?S - e"
+        by (intro cSUP_least[OF D(1)]) auto
       then show False
         using goal2 by auto
     qed
@@ -11813,9 +11753,9 @@
     proof -
       fix a b :: 'n
       assume ab: "ball 0 (K + 1) \<subseteq> {a..b}"
-      have *: "\<forall>s s1. i - e < s1 \<and> s1 \<le> s \<and> s < i + e \<longrightarrow> abs (s - i) < e"
+      have *: "\<forall>s s1. ?S - e < s1 \<and> s1 \<le> s \<and> s < ?S + e \<longrightarrow> abs (s - ?S) < e"
         by arith
-      show "norm (integral {a..b} (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) - i) < e"
+      show "norm (integral {a..b} (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) - ?S) < e"
         unfolding real_norm_def
         apply (rule *[rule_format])
         apply safe
@@ -11867,10 +11807,10 @@
         from henstock_lemma[OF f(1) `e/2>0`] guess d2 . note d2=this
         from fine_division_exists[OF gauge_inter[OF d1(1) d2(1)], of a b] guess p .
         note p=this(1) conjunctD2[OF this(2)[unfolded fine_inter]]
-        have *: "\<And>sf sf' si di. sf' = sf \<longrightarrow> si \<le> i \<longrightarrow> abs (sf - si) < e / 2 \<longrightarrow>
-          abs (sf' - di) < e / 2 \<longrightarrow> di < i + e"
+        have *: "\<And>sf sf' si di. sf' = sf \<longrightarrow> si \<le> ?S \<longrightarrow> abs (sf - si) < e / 2 \<longrightarrow>
+          abs (sf' - di) < e / 2 \<longrightarrow> di < ?S + e"
           by arith
-        show "integral {a..b} (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) < i + e"
+        show "integral {a..b} (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) < ?S + e"
           apply (subst if_P)
           apply rule
         proof (rule *[rule_format])
@@ -11893,18 +11833,12 @@
             apply (subst abs_of_nonneg)
             apply auto
             done
-          show "(\<Sum>(x, k)\<in>p. norm (integral k f)) \<le> i"
+          show "(\<Sum>(x, k)\<in>p. norm (integral k f)) \<le> ?S"
+            using partial_division_of_tagged_division[of p "{a..b}"] p(1)
             apply (subst setsum_over_tagged_division_lemma[OF p(1)])
-            defer
-            apply (rule isLubD2[OF i])
-            unfolding image_iff
-            apply (rule_tac x="snd ` p" in bexI)
-            unfolding mem_Collect_eq
-            defer
-            apply (rule partial_division_of_tagged_division[of _ "{a..b}"])
-            using p(1)
-            unfolding tagged_division_of_def
-            apply auto
+            apply (simp add: integral_null)
+            apply (intro cSUP_upper2[OF D(2), of "snd ` p"])
+            apply (auto simp: tagged_partial_division_of_def)
             done
         qed
       qed
@@ -11975,7 +11909,7 @@
     and "g absolutely_integrable_on s"
   shows "(\<lambda>x. f x - g x) absolutely_integrable_on s"
   using absolutely_integrable_add[OF assms(1) absolutely_integrable_neg[OF assms(2)]]
-  by (simp only: algebra_simps)
+  by (simp add: algebra_simps)
 
 lemma absolutely_integrable_linear:
   fixes f :: "'m::ordered_euclidean_space \<Rightarrow> 'n::ordered_euclidean_space"
@@ -12380,11 +12314,22 @@
 lemma dominated_convergence:
   fixes f :: "nat \<Rightarrow> 'n::ordered_euclidean_space \<Rightarrow> real"
   assumes "\<And>k. (f k) integrable_on s" "h integrable_on s"
-    and "\<And>k. \<forall>x \<in> s. norm(f k x) \<le> (h x)"
+    and "\<And>k. \<forall>x \<in> s. norm (f k x) \<le> h x"
     and "\<forall>x \<in> s. ((\<lambda>k. f k x) ---> g x) sequentially"
   shows "g integrable_on s"
     and "((\<lambda>k. integral s (f k)) ---> integral s g) sequentially"
 proof -
+  have bdd_below[simp]: "\<And>x P. x \<in> s \<Longrightarrow> bdd_below {f n x |n. P n}"
+  proof (safe intro!: bdd_belowI)
+    fix n x show "x \<in> s \<Longrightarrow> - h x \<le> f n x"
+      using assms(3)[rule_format, of x n] by simp
+  qed
+  have bdd_above[simp]: "\<And>x P. x \<in> s \<Longrightarrow> bdd_above {f n x |n. P n}"
+  proof (safe intro!: bdd_aboveI)
+    fix n x show "x \<in> s \<Longrightarrow> f n x \<le> h x"
+      using assms(3)[rule_format, of x n] by simp
+  qed
+
   have "\<And>m. (\<lambda>x. Inf {f j x |j. m \<le> j}) integrable_on s \<and>
     ((\<lambda>k. integral s (\<lambda>x. Inf {f j x |j. j \<in> {m..m + k}})) --->
     integral s (\<lambda>x. Inf {f j x |j. m \<le> j}))sequentially"
@@ -12424,66 +12369,32 @@
     fix x
     assume x: "x \<in> s"
     show "Inf {f j x |j. j \<in> {m..m + Suc k}} \<le> Inf {f j x |j. j \<in> {m..m + k}}"
-      apply (rule cInf_ge)
-      unfolding setge_def
-      defer
-      apply rule
-      apply (subst cInf_finite_le_iff)
-      prefer 3
-      apply (rule_tac x=xa in bexI)
-      apply auto
-      done
-    let ?S = "{f j x| j.  m \<le> j}"
-    def i \<equiv> "Inf ?S"
-    show "((\<lambda>k. Inf {f j x |j. j \<in> {m..m + k}}) ---> i) sequentially"
+      by (rule cInf_superset_mono) auto
+    let ?S = "{f j x| j. m \<le> j}"
+    show "((\<lambda>k. Inf {f j x |j. j \<in> {m..m + k}}) ---> Inf ?S) sequentially"
     proof (rule LIMSEQ_I)
       case goal1
       note r = this
-      have i: "isGlb UNIV ?S i"
-        unfolding i_def
-        apply (rule Inf)
-        defer
-        apply (rule_tac x="- h x - 1" in exI)
-        unfolding setge_def
-      proof safe
-        case goal1
-        then show ?case using assms(3)[rule_format,OF x, of j] by auto
-      qed auto
-
-      have "\<exists>y\<in>?S. \<not> y \<ge> i + r"
-      proof(rule ccontr)
-        assume "\<not> ?thesis"
-        then have "i \<ge> i + r"
-          apply -
-          apply (rule isGlb_le_isLb[OF i])
-          apply (rule isLbI)
-          unfolding setge_def
-          apply fastforce+
-          done
-        then show False using r by auto
-      qed
-      then guess y .. note y=this[unfolded not_le]
-      from this(1)[unfolded mem_Collect_eq] guess N .. note N=conjunctD2[OF this]
+
+      have "\<exists>y\<in>?S. y < Inf ?S + r"
+        by (subst cInf_less_iff[symmetric]) (auto simp: `x\<in>s` r)
+      then obtain N where N: "f N x < Inf ?S + r" "m \<le> N"
+        by blast
 
       show ?case
         apply (rule_tac x=N in exI)
       proof safe
         case goal1
-        have *: "\<And>y ix. y < i + r \<longrightarrow> i \<le> ix \<longrightarrow> ix \<le> y \<longrightarrow> abs(ix - i) < r"
+        have *: "\<And>y ix. y < Inf ?S + r \<longrightarrow> Inf ?S \<le> ix \<longrightarrow> ix \<le> y \<longrightarrow> abs(ix - Inf ?S) < r"
           by arith
         show ?case
           unfolding real_norm_def
-            apply (rule *[rule_format,OF y(2)])
-            unfolding i_def
-            apply (rule real_le_inf_subset)
-            prefer 3
-            apply (rule,rule isGlbD1[OF i])
-            prefer 3
-            apply (subst cInf_finite_le_iff)
-            prefer 3
-            apply (rule_tac x=y in bexI)
+            apply (rule *[rule_format, OF N(1)])
+            apply (rule cInf_superset_mono, auto simp: `x\<in>s`) []
+            apply (rule cInf_lower)
             using N goal1
-            apply auto
+            apply auto []
+            apply simp
             done
       qed
     qed
@@ -12527,65 +12438,27 @@
     fix x
     assume x: "x\<in>s"
     show "Sup {f j x |j. j \<in> {m..m + Suc k}} \<ge> Sup {f j x |j. j \<in> {m..m + k}}"
-      apply (rule cSup_le)
-      unfolding setle_def
-      defer
-      apply rule
-      apply (subst cSup_finite_ge_iff)
-      prefer 3
-      apply (rule_tac x=y in bexI)
-      apply auto
-      done
-    let ?S = "{f j x| j.  m \<le> j}"
-    def i \<equiv> "Sup ?S"
-    show "((\<lambda>k. Sup {f j x |j. j \<in> {m..m + k}}) ---> i) sequentially"
+      by (rule cSup_subset_mono) auto
+    let ?S = "{f j x| j. m \<le> j}"
+    show "((\<lambda>k. Sup {f j x |j. j \<in> {m..m + k}}) ---> Sup ?S) sequentially"
     proof (rule LIMSEQ_I)
       case goal1 note r=this
-      have i: "isLub UNIV ?S i"
-        unfolding i_def
-        apply (rule isLub_cSup)
-        defer
-        apply (rule_tac x="h x" in exI)
-        unfolding setle_def
-      proof safe
-        case goal1
-        then show ?case
-          using assms(3)[rule_format,OF x, of j] by auto
-      qed auto
-
-      have "\<exists>y\<in>?S. \<not> y \<le> i - r"
-      proof (rule ccontr)
-        assume "\<not> ?thesis"
-        then have "i \<le> i - r"
-          apply -
-          apply (rule isLub_le_isUb[OF i])
-          apply (rule isUbI)
-          unfolding setle_def
-          apply fastforce+
-          done
-        then show False
-          using r by auto
-      qed
-      then guess y .. note y=this[unfolded not_le]
-      from this(1)[unfolded mem_Collect_eq] guess N .. note N=conjunctD2[OF this]
+      have "\<exists>y\<in>?S. Sup ?S - r < y"
+        by (subst less_cSup_iff[symmetric]) (auto simp: r `x\<in>s`)
+      then obtain N where N: "Sup ?S - r < f N x" "m \<le> N"
+        by blast
 
       show ?case
         apply (rule_tac x=N in exI)
       proof safe
         case goal1
-        have *: "\<And>y ix. i - r < y \<longrightarrow> ix \<le> i \<longrightarrow> y \<le> ix \<longrightarrow> abs(ix - i) < r"
+        have *: "\<And>y ix. Sup ?S - r < y \<longrightarrow> ix \<le> Sup ?S \<longrightarrow> y \<le> ix \<longrightarrow> abs(ix - Sup ?S) < r"
           by arith
         show ?case
-          unfolding real_norm_def
-          apply (rule *[rule_format,OF y(2)])
-          unfolding i_def
-          apply (rule real_ge_sup_subset)
-          prefer 3
-          apply (rule, rule isLubD1[OF i])
-          prefer 3
-          apply (subst cSup_finite_ge_iff)
-          prefer 3
-          apply (rule_tac x = y in bexI)
+          apply simp
+          apply (rule *[rule_format, OF N(1)])
+          apply (rule cSup_subset_mono, auto simp: `x\<in>s`) []
+          apply (subst cSup_upper)
           using N goal1
           apply auto
           done
@@ -12618,17 +12491,7 @@
 
     have *: "\<And>x y::real. x \<ge> - y \<Longrightarrow> - x \<le> y" by auto
     show "Inf {f j x |j. k \<le> j} \<le> Inf {f j x |j. Suc k \<le> j}"
-      apply -
-      apply (rule real_le_inf_subset)
-      prefer 3
-      unfolding setge_def
-      apply (rule_tac x="- h x" in exI)
-      apply safe
-      apply (rule *)
-      using assms(3)[rule_format,OF x]
-      unfolding real_norm_def abs_le_iff
-      apply auto
-      done
+      by (intro cInf_superset_mono) (auto simp: `x\<in>s`)
 
     show "(\<lambda>k::nat. Inf {f j x |j. k \<le> j}) ----> g x"
     proof (rule LIMSEQ_I)
@@ -12676,16 +12539,7 @@
     assume x: "x \<in> s"
 
     show "Sup {f j x |j. k \<le> j} \<ge> Sup {f j x |j. Suc k \<le> j}"
-      apply -
-      apply (rule real_ge_sup_subset)
-      prefer 3
-      unfolding setle_def
-      apply (rule_tac x="h x" in exI)
-      apply safe
-      using assms(3)[rule_format,OF x]
-      unfolding real_norm_def abs_le_iff
-      apply auto
-      done
+      by (rule cSup_subset_mono) (auto simp: `x\<in>s`)
     show "((\<lambda>k. Sup {f j x |j. k \<le> j}) ---> g x) sequentially"
     proof (rule LIMSEQ_I)
       case goal1
@@ -12714,42 +12568,18 @@
     from LIMSEQ_D [OF inc2(2) goal1] guess N1 .. note N1=this[unfolded real_norm_def]
     from LIMSEQ_D [OF dec2(2) goal1] guess N2 .. note N2=this[unfolded real_norm_def]
     show ?case
-      apply (rule_tac x="N1+N2" in exI, safe)
-    proof -
+    proof (rule_tac x="N1+N2" in exI, safe)
       fix n
       assume n: "n \<ge> N1 + N2"
       have *: "\<And>i0 i i1 g. \<bar>i0 - g\<bar> < r \<longrightarrow> \<bar>i1 - g\<bar> < r \<longrightarrow> i0 \<le> i \<longrightarrow> i \<le> i1 \<longrightarrow> \<bar>i - g\<bar> < r"
         by arith
       show "norm (integral s (f n) - integral s g) < r"
         unfolding real_norm_def
-        apply (rule *[rule_format,OF N1[rule_format] N2[rule_format], of n n])
-      proof -
+      proof (rule *[rule_format,OF N1[rule_format] N2[rule_format], of n n])
         show "integral s (\<lambda>x. Inf {f j x |j. n \<le> j}) \<le> integral s (f n)"
-        proof (rule integral_le[OF dec1(1) assms(1)], safe)
-          fix x
-          assume x: "x \<in> s"
-          have *: "\<And>x y::real. x \<ge> - y \<Longrightarrow> - x \<le> y" by auto
-          show "Inf {f j x |j. n \<le> j} \<le> f n x"
-            apply (rule cInf_lower[where z="- h x"])
-            defer
-            apply (rule *)
-            using assms(3)[rule_format,OF x]
-            unfolding real_norm_def abs_le_iff
-            apply auto
-            done
-        qed
+          by (rule integral_le[OF dec1(1) assms(1)]) (auto intro!: cInf_lower)
         show "integral s (f n) \<le> integral s (\<lambda>x. Sup {f j x |j. n \<le> j})"
-        proof (rule integral_le[OF assms(1) inc1(1)], safe)
-          fix x
-          assume x: "x \<in> s"
-          show "f n x \<le> Sup {f j x |j. n \<le> j}"
-            apply (rule cSup_upper[where z="h x"])
-            defer
-            using assms(3)[rule_format,OF x]
-            unfolding real_norm_def abs_le_iff
-            apply auto
-            done
-        qed
+          by (rule integral_le[OF assms(1) inc1(1)]) (auto intro!: cSup_upper)
       qed (insert n, auto)
     qed
   qed
--- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -303,7 +303,7 @@
   by (metis linear_iff)
 
 lemma linear_sub: "linear f \<Longrightarrow> f (x - y) = f x - f y"
-  by (simp add: diff_minus linear_add linear_neg)
+  using linear_add [of f x "- y"] by (simp add: linear_neg)
 
 lemma linear_setsum:
   assumes lin: "linear f"
@@ -384,10 +384,10 @@
   using bilinear_radd [OF assms, of x 0 0 ] by (simp add: eq_add_iff field_simps)
 
 lemma bilinear_lsub: "bilinear h \<Longrightarrow> h (x - y) z = h x z - h y z"
-  by (simp  add: diff_minus bilinear_ladd bilinear_lneg)
+  using bilinear_ladd [of h x "- y"] by (simp add: bilinear_lneg)
 
 lemma bilinear_rsub: "bilinear h \<Longrightarrow> h z (x - y) = h z x - h z y"
-  by (simp  add: diff_minus bilinear_radd bilinear_rneg)
+  using bilinear_radd [of h _ x "- y"] by (simp add: bilinear_rneg)
 
 lemma bilinear_setsum:
   assumes bh: "bilinear h"
@@ -730,7 +730,7 @@
   by (metis scaleR_minus1_left subspace_mul)
 
 lemma subspace_sub: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x - y \<in> S"
-  by (metis diff_minus subspace_add subspace_neg)
+  using subspace_add [of S x "- y"] by (simp add: subspace_neg)
 
 lemma (in real_vector) subspace_setsum:
   assumes sA: "subspace A"
@@ -1021,8 +1021,7 @@
     apply safe
     apply (rule_tac x=k in exI, simp)
     apply (erule rev_image_eqI [OF SigmaI [OF rangeI]])
-    apply simp
-    apply (rule right_minus)
+    apply auto
     done
   then show ?thesis by simp
 qed
@@ -2064,7 +2063,7 @@
       using C by simp
     have "orthogonal ?a y"
       unfolding orthogonal_def
-      unfolding inner_diff inner_setsum_left diff_eq_0_iff_eq
+      unfolding inner_diff inner_setsum_left right_minus_eq
       unfolding setsum_diff1' [OF `finite C` `y \<in> C`]
       apply (clarsimp simp add: inner_commute[of y a])
       apply (rule setsum_0')
@@ -3026,7 +3025,7 @@
         norm x * (norm y * (y \<bullet> x) - norm x * norm y * norm y) =  0)"
       using x y
       unfolding inner_simps
-      unfolding power2_norm_eq_inner[symmetric] power2_eq_square diff_eq_0_iff_eq
+      unfolding power2_norm_eq_inner[symmetric] power2_eq_square right_minus_eq
       apply (simp add: inner_commute)
       apply (simp add: field_simps)
       apply metis
--- a/src/HOL/Multivariate_Analysis/Operator_Norm.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Operator_Norm.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -8,7 +8,7 @@
 imports Linear_Algebra
 begin
 
-definition "onorm f = Sup {norm (f x)| x. norm x = 1}"
+definition "onorm f = (SUP x:{x. norm x = 1}. norm (f x))"
 
 lemma norm_bound_generalize:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
@@ -67,25 +67,22 @@
   shows "norm (f x) \<le> onorm f * norm x"
     and "\<forall>x. norm (f x) \<le> b * norm x \<Longrightarrow> onorm f \<le> b"
 proof -
-  let ?S = "{norm (f x) |x. norm x = 1}"
+  let ?S = "(\<lambda>x. norm (f x))`{x. norm x = 1}"
   have "norm (f (SOME i. i \<in> Basis)) \<in> ?S"
     by (auto intro!: exI[of _ "SOME i. i \<in> Basis"] norm_Basis SOME_Basis)
   then have Se: "?S \<noteq> {}"
     by auto
-  from linear_bounded[OF lf] have b: "\<exists> b. ?S *<= b"
-    unfolding norm_bound_generalize[OF lf, symmetric]
-    by (auto simp add: setle_def)
-  from isLub_cSup[OF Se b, unfolded onorm_def[symmetric]]
-  show "norm (f x) \<le> onorm f * norm x"
+  from linear_bounded[OF lf] have b: "bdd_above ?S"
+    unfolding norm_bound_generalize[OF lf, symmetric] by auto
+  then show "norm (f x) \<le> onorm f * norm x"
     apply -
     apply (rule spec[where x = x])
     unfolding norm_bound_generalize[OF lf, symmetric]
-    apply (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)
+    apply (auto simp: onorm_def intro!: cSUP_upper)
     done
   show "\<forall>x. norm (f x) \<le> b * norm x \<Longrightarrow> onorm f \<le> b"
-    using isLub_cSup[OF Se b, unfolded onorm_def[symmetric]]
     unfolding norm_bound_generalize[OF lf, symmetric]
-    by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)
+    using Se by (auto simp: onorm_def intro!: cSUP_least b)
 qed
 
 lemma onorm_pos_le:
@@ -107,18 +104,8 @@
   apply arith
   done
 
-lemma onorm_const:
-  "onorm (\<lambda>x::'a::euclidean_space. y::'b::euclidean_space) = norm y"
-proof -
-  let ?f = "\<lambda>x::'a. y::'b"
-  have th: "{norm (?f x)| x. norm x = 1} = {norm y}"
-    by (auto simp: SOME_Basis intro!: exI[of _ "SOME i. i \<in> Basis"])
-  show ?thesis
-    unfolding onorm_def th
-    apply (rule cSup_unique)
-    apply (simp_all  add: setle_def)
-    done
-qed
+lemma onorm_const: "onorm (\<lambda>x::'a::euclidean_space. y::'b::euclidean_space) = norm y"
+  using SOME_Basis by (auto simp add: onorm_def intro!: cSUP_const norm_Basis)
 
 lemma onorm_pos_lt:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -10,7 +10,6 @@
 imports
   Complex_Main
   "~~/src/HOL/Library/Countable_Set"
-  "~~/src/HOL/Library/Glbs"
   "~~/src/HOL/Library/FuncSet"
   Linear_Algebra
   Norm_Arith
@@ -28,8 +27,6 @@
 lemma lim_subseq: "subseq r \<Longrightarrow> s ----> l \<Longrightarrow> (s \<circ> r) ----> l"
   by (rule LIMSEQ_subseq_LIMSEQ)
 
-lemmas real_isGlb_unique = isGlb_unique[where 'a=real]
-
 lemma countable_PiE:
   "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> countable (F i)) \<Longrightarrow> countable (PiE I F)"
   by (induct I arbitrary: F rule: finite_induct) (auto simp: PiE_insert_eq)
@@ -789,18 +786,20 @@
   "a - b \<ge> c \<longleftrightarrow> a \<ge> c + b"
   by arith+
 
-lemma open_ball[intro, simp]: "open (ball x e)"
-  unfolding open_dist ball_def mem_Collect_eq Ball_def
-  unfolding dist_commute
-  apply clarify
-  apply (rule_tac x="e - dist xa x" in exI)
-  using dist_triangle_alt[where z=x]
-  apply (clarsimp simp add: diff_less_iff)
-  apply atomize
-  apply (erule_tac x="y" in allE)
-  apply (erule_tac x="xa" in allE)
-  apply arith
-  done
+lemma open_vimage: (* TODO: move to Topological_Spaces.thy *)
+  assumes "open s" and "continuous_on UNIV f"
+  shows "open (vimage f s)"
+  using assms unfolding continuous_on_open_vimage [OF open_UNIV]
+  by simp
+
+lemma open_ball [intro, simp]: "open (ball x e)"
+proof -
+  have "open (dist x -` {..<e})"
+    by (intro open_vimage open_lessThan continuous_on_intros)
+  also have "dist x -` {..<e} = ball x e"
+    by auto
+  finally show ?thesis .
+qed
 
 lemma open_contains_ball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. ball x e \<subseteq> S)"
   unfolding open_dist subset_eq mem_ball Ball_def dist_commute ..
@@ -1553,7 +1552,7 @@
       fix y
       assume "y \<in> {x<..} \<inter> I"
       with False bnd have "Inf (f ` ({x<..} \<inter> I)) \<le> f y"
-        by (auto intro: cInf_lower)
+        by (auto intro!: cInf_lower bdd_belowI2)
       with a have "a < f y"
         by (blast intro: less_le_trans)
     }
@@ -1889,7 +1888,6 @@
 lemma closed_sequential_limits:
   fixes S :: "'a::first_countable_topology set"
   shows "closed S \<longleftrightarrow> (\<forall>x l. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially \<longrightarrow> l \<in> S)"
-  unfolding closed_limpt
   using closure_sequential [where 'a='a] closure_closed [where 'a='a]
     closed_limpt [where 'a='a] islimpt_sequential [where 'a='a] mem_delete [where 'a='a]
   by metis
@@ -1908,17 +1906,17 @@
 
 lemma closure_contains_Inf:
   fixes S :: "real set"
-  assumes "S \<noteq> {}" "\<forall>x\<in>S. B \<le> x"
+  assumes "S \<noteq> {}" "bdd_below S"
   shows "Inf S \<in> closure S"
 proof -
   have *: "\<forall>x\<in>S. Inf S \<le> x"
-    using cInf_lower_EX[of _ S] assms by metis
+    using cInf_lower[of _ S] assms by metis
   {
     fix e :: real
     assume "e > 0"
     then have "Inf S < Inf S + e" by simp
     with assms obtain x where "x \<in> S" "x < Inf S + e"
-      by (subst (asm) cInf_less_iff[of _ B]) auto
+      by (subst (asm) cInf_less_iff) auto
     with * have "\<exists>x\<in>S. dist x (Inf S) < e"
       by (intro bexI[of _ x]) (auto simp add: dist_real_def)
   }
@@ -1927,12 +1925,9 @@
 
 lemma closed_contains_Inf:
   fixes S :: "real set"
-  assumes "S \<noteq> {}" "\<forall>x\<in>S. B \<le> x"
-    and "closed S"
-  shows "Inf S \<in> S"
+  shows "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> closed S \<Longrightarrow> Inf S \<in> S"
   by (metis closure_contains_Inf closure_closed assms)
 
-
 lemma not_trivial_limit_within_ball:
   "\<not> trivial_limit (at x within S) \<longleftrightarrow> (\<forall>e>0. S \<inter> ball x e - {x} \<noteq> {})"
   (is "?lhs = ?rhs")
@@ -1974,29 +1969,25 @@
 
 subsection {* Infimum Distance *}
 
-definition "infdist x A = (if A = {} then 0 else Inf {dist x a|a. a \<in> A})"
-
-lemma infdist_notempty: "A \<noteq> {} \<Longrightarrow> infdist x A = Inf {dist x a|a. a \<in> A}"
+definition "infdist x A = (if A = {} then 0 else INF a:A. dist x a)"
+
+lemma bdd_below_infdist[intro, simp]: "bdd_below (dist x`A)"
+  by (auto intro!: zero_le_dist)
+
+lemma infdist_notempty: "A \<noteq> {} \<Longrightarrow> infdist x A = (INF a:A. dist x a)"
   by (simp add: infdist_def)
 
 lemma infdist_nonneg: "0 \<le> infdist x A"
-  by (auto simp add: infdist_def intro: cInf_greatest)
-
-lemma infdist_le:
-  assumes "a \<in> A"
-    and "d = dist x a"
-  shows "infdist x A \<le> d"
-  using assms by (auto intro!: cInf_lower[where z=0] simp add: infdist_def)
-
-lemma infdist_zero[simp]:
-  assumes "a \<in> A"
-  shows "infdist a A = 0"
-proof -
-  from infdist_le[OF assms, of "dist a a"] have "infdist a A \<le> 0"
-    by auto
-  with infdist_nonneg[of a A] assms show "infdist a A = 0"
-    by auto
-qed
+  by (auto simp add: infdist_def intro: cINF_greatest)
+
+lemma infdist_le: "a \<in> A \<Longrightarrow> infdist x A \<le> dist x a"
+  by (auto intro: cINF_lower simp add: infdist_def)
+
+lemma infdist_le2: "a \<in> A \<Longrightarrow> dist x a \<le> d \<Longrightarrow> infdist x A \<le> d"
+  by (auto intro!: cINF_lower2 simp add: infdist_def)
+
+lemma infdist_zero[simp]: "a \<in> A \<Longrightarrow> infdist a A = 0"
+  by (auto intro!: antisym infdist_nonneg infdist_le2)
 
 lemma infdist_triangle: "infdist x A \<le> infdist y A + dist x y"
 proof (cases "A = {}")
@@ -2015,18 +2006,11 @@
       by auto
     show "infdist x A \<le> d"
       unfolding infdist_notempty[OF `A \<noteq> {}`]
-    proof (rule cInf_lower2)
-      show "dist x a \<in> {dist x a |a. a \<in> A}"
-        using `a \<in> A` by auto
+    proof (rule cINF_lower2)
+      show "a \<in> A" by fact
       show "dist x a \<le> d"
         unfolding d by (rule dist_triangle)
-      fix d
-      assume "d \<in> {dist x a |a. a \<in> A}"
-      then obtain a where "a \<in> A" "d = dist x a"
-        by auto
-      then show "infdist x A \<le> d"
-        by (rule infdist_le)
-    qed
+    qed simp
   qed
   also have "\<dots> = dist x y + infdist y A"
   proof (rule cInf_eq, safe)
@@ -2039,7 +2023,7 @@
     assume inf: "\<And>d. d \<in> {dist x y + dist y a |a. a \<in> A} \<Longrightarrow> i \<le> d"
     then have "i - dist x y \<le> infdist y A"
       unfolding infdist_notempty[OF `A \<noteq> {}`] using `a \<in> A`
-      by (intro cInf_greatest) (auto simp: field_simps)
+      by (intro cINF_greatest) (auto simp: field_simps)
     then show "i \<le> dist x y + infdist y A"
       by simp
   qed
@@ -2078,7 +2062,7 @@
     assume "\<not> (\<exists>y\<in>A. dist y x < e)"
     then have "infdist x A \<ge> e" using `a \<in> A`
       unfolding infdist_def
-      by (force simp: dist_commute intro: cInf_greatest)
+      by (force simp: dist_commute intro: cINF_greatest)
     with x `e > 0` show False by auto
   qed
 qed
@@ -2129,15 +2113,20 @@
 
 subsection {* More properties of closed balls *}
 
+lemma closed_vimage: (* TODO: move to Topological_Spaces.thy *)
+  assumes "closed s" and "continuous_on UNIV f"
+  shows "closed (vimage f s)"
+  using assms unfolding continuous_on_closed_vimage [OF closed_UNIV]
+  by simp
+
 lemma closed_cball: "closed (cball x e)"
-  unfolding cball_def closed_def
-  unfolding Collect_neg_eq [symmetric] not_le
-  apply (clarsimp simp add: open_dist, rename_tac y)
-  apply (rule_tac x="dist x y - e" in exI, clarsimp)
-  apply (rename_tac x')
-  apply (cut_tac x=x and y=x' and z=y in dist_triangle)
-  apply simp
-  done
+proof -
+  have "closed (dist x -` {..e})"
+    by (intro closed_vimage closed_atMost continuous_on_intros)
+  also have "dist x -` {..e} = cball x e"
+    by auto
+  finally show ?thesis .
+qed
 
 lemma open_contains_cball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0.  cball x e \<subseteq> S)"
 proof -
@@ -2645,11 +2634,19 @@
 
 text{* Some theorems on sups and infs using the notion "bounded". *}
 
-lemma bounded_real:
-  fixes S :: "real set"
-  shows "bounded S \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. abs x \<le> a)"
+lemma bounded_real: "bounded (S::real set) \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. \<bar>x\<bar> \<le> a)"
   by (simp add: bounded_iff)
 
+lemma bounded_imp_bdd_above: "bounded S \<Longrightarrow> bdd_above (S :: real set)"
+  by (auto simp: bounded_def bdd_above_def dist_real_def)
+     (metis abs_le_D1 abs_minus_commute diff_le_eq)
+
+lemma bounded_imp_bdd_below: "bounded S \<Longrightarrow> bdd_below (S :: real set)"
+  by (auto simp: bounded_def bdd_below_def dist_real_def)
+     (metis abs_le_D1 add_commute diff_le_eq)
+
+(* TODO: remove the following lemmas about Inf and Sup, is now in conditionally complete lattice *)
+
 lemma bounded_has_Sup:
   fixes S :: "real set"
   assumes "bounded S"
@@ -2657,22 +2654,14 @@
   shows "\<forall>x\<in>S. x \<le> Sup S"
     and "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b"
 proof
-  fix x
-  assume "x\<in>S"
-  then show "x \<le> Sup S"
-    by (metis cSup_upper abs_le_D1 assms(1) bounded_real)
-next
   show "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b"
     using assms by (metis cSup_least)
-qed
+qed (metis cSup_upper assms(1) bounded_imp_bdd_above)
 
 lemma Sup_insert:
   fixes S :: "real set"
   shows "bounded S \<Longrightarrow> Sup (insert x S) = (if S = {} then x else max x (Sup S))"
-  apply (subst cSup_insert_If)
-  apply (rule bounded_has_Sup(1)[of S, rule_format])
-  apply (auto simp: sup_max)
-  done
+  by (auto simp: bounded_imp_bdd_above sup_max cSup_insert_If)
 
 lemma Sup_insert_finite:
   fixes S :: "real set"
@@ -2689,24 +2678,14 @@
   shows "\<forall>x\<in>S. x \<ge> Inf S"
     and "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b"
 proof
-  fix x
-  assume "x \<in> S"
-  from assms(1) obtain a where a: "\<forall>x\<in>S. \<bar>x\<bar> \<le> a"
-    unfolding bounded_real by auto
-  then show "x \<ge> Inf S" using `x \<in> S`
-    by (metis cInf_lower_EX abs_le_D2 minus_le_iff)
-next
   show "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b"
     using assms by (metis cInf_greatest)
-qed
+qed (metis cInf_lower assms(1) bounded_imp_bdd_below)
 
 lemma Inf_insert:
   fixes S :: "real set"
   shows "bounded S \<Longrightarrow> Inf (insert x S) = (if S = {} then x else min x (Inf S))"
-  apply (subst cInf_insert_if)
-  apply (rule bounded_has_Inf(1)[of S, rule_format])
-  apply (auto simp: inf_min)
-  done
+  by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_If)
 
 lemma Inf_insert_finite:
   fixes S :: "real set"
@@ -3298,6 +3277,50 @@
   where "seq_compact S \<longleftrightarrow>
     (\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially))"
 
+lemma seq_compactI:
+  assumes "\<And>f. \<forall>n. f n \<in> S \<Longrightarrow> \<exists>l\<in>S. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
+  shows "seq_compact S"
+  unfolding seq_compact_def using assms by fast
+
+lemma seq_compactE:
+  assumes "seq_compact S" "\<forall>n. f n \<in> S"
+  obtains l r where "l \<in> S" "subseq r" "((f \<circ> r) ---> l) sequentially"
+  using assms unfolding seq_compact_def by fast
+
+lemma closed_sequentially: (* TODO: move upwards *)
+  assumes "closed s" and "\<forall>n. f n \<in> s" and "f ----> l"
+  shows "l \<in> s"
+proof (rule ccontr)
+  assume "l \<notin> s"
+  with `closed s` and `f ----> l` have "eventually (\<lambda>n. f n \<in> - s) sequentially"
+    by (fast intro: topological_tendstoD)
+  with `\<forall>n. f n \<in> s` show "False"
+    by simp
+qed
+
+lemma seq_compact_inter_closed:
+  assumes "seq_compact s" and "closed t"
+  shows "seq_compact (s \<inter> t)"
+proof (rule seq_compactI)
+  fix f assume "\<forall>n::nat. f n \<in> s \<inter> t"
+  hence "\<forall>n. f n \<in> s" and "\<forall>n. f n \<in> t"
+    by simp_all
+  from `seq_compact s` and `\<forall>n. f n \<in> s`
+  obtain l r where "l \<in> s" and r: "subseq r" and l: "(f \<circ> r) ----> l"
+    by (rule seq_compactE)
+  from `\<forall>n. f n \<in> t` have "\<forall>n. (f \<circ> r) n \<in> t"
+    by simp
+  from `closed t` and this and l have "l \<in> t"
+    by (rule closed_sequentially)
+  with `l \<in> s` and r and l show "\<exists>l\<in>s \<inter> t. \<exists>r. subseq r \<and> (f \<circ> r) ----> l"
+    by fast
+qed
+
+lemma seq_compact_closed_subset:
+  assumes "closed s" and "s \<subseteq> t" and "seq_compact t"
+  shows "seq_compact s"
+  using assms seq_compact_inter_closed [of t s] by (simp add: Int_absorb1)
+
 lemma seq_compact_imp_countably_compact:
   fixes U :: "'a :: first_countable_topology set"
   assumes "seq_compact U"
@@ -3410,16 +3433,6 @@
     using `x \<in> U` by (auto simp: convergent_def comp_def)
 qed
 
-lemma seq_compactI:
-  assumes "\<And>f. \<forall>n. f n \<in> S \<Longrightarrow> \<exists>l\<in>S. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
-  shows "seq_compact S"
-  unfolding seq_compact_def using assms by fast
-
-lemma seq_compactE:
-  assumes "seq_compact S" "\<forall>n. f n \<in> S"
-  obtains l r where "l \<in> S" "subseq r" "((f \<circ> r) ---> l) sequentially"
-  using assms unfolding seq_compact_def by fast
-
 lemma countably_compact_imp_acc_point:
   assumes "countably_compact s"
     and "countable t"
@@ -3654,6 +3667,8 @@
   "\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t) \<Longrightarrow> bounded s"
   using compact_imp_bounded unfolding compact_eq_bolzano_weierstrass .
 
+subsection {* Metric spaces with the Heine-Borel property *}
+
 text {*
   A metric space (or topological vector space) is said to have the
   Heine-Borel property if every closed and bounded subset is compact.
@@ -3678,7 +3693,7 @@
   from f have fr: "\<forall>n. (f \<circ> r) n \<in> s"
     by simp
   have "l \<in> s" using `closed s` fr l
-    unfolding closed_sequential_limits by blast
+    by (rule closed_sequentially)
   show "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
     using `l \<in> s` r l by blast
 qed
@@ -3859,11 +3874,21 @@
     using l r by fast
 qed
 
-subsubsection{* Completeness *}
+subsubsection {* Completeness *}
 
 definition complete :: "'a::metric_space set \<Rightarrow> bool"
   where "complete s \<longleftrightarrow> (\<forall>f. (\<forall>n. f n \<in> s) \<and> Cauchy f \<longrightarrow> (\<exists>l\<in>s. f ----> l))"
 
+lemma completeI:
+  assumes "\<And>f. \<forall>n. f n \<in> s \<Longrightarrow> Cauchy f \<Longrightarrow> \<exists>l\<in>s. f ----> l"
+  shows "complete s"
+  using assms unfolding complete_def by fast
+
+lemma completeE:
+  assumes "complete s" and "\<forall>n. f n \<in> s" and "Cauchy f"
+  obtains l where "l \<in> s" and "f ----> l"
+  using assms unfolding complete_def by fast
+
 lemma compact_imp_complete:
   assumes "compact s"
   shows "complete s"
@@ -4085,49 +4110,57 @@
 
 instance euclidean_space \<subseteq> banach ..
 
-lemma complete_univ: "complete (UNIV :: 'a::complete_space set)"
-proof (simp add: complete_def, rule, rule)
-  fix f :: "nat \<Rightarrow> 'a"
-  assume "Cauchy f"
+lemma complete_UNIV: "complete (UNIV :: ('a::complete_space) set)"
+proof (rule completeI)
+  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
   then have "convergent f" by (rule Cauchy_convergent)
-  then show "\<exists>l. f ----> l" unfolding convergent_def .
+  then show "\<exists>l\<in>UNIV. f ----> l" unfolding convergent_def by simp
 qed
 
 lemma complete_imp_closed:
   assumes "complete s"
   shows "closed s"
-proof -
-  {
-    fix x
-    assume "x islimpt s"
-    then obtain f where f: "\<forall>n. f n \<in> s - {x}" "(f ---> x) sequentially"
-      unfolding islimpt_sequential by auto
-    then obtain l where l: "l\<in>s" "(f ---> l) sequentially"
-      using `complete s`[unfolded complete_def] using LIMSEQ_imp_Cauchy[of f x] by auto
-    then have "x \<in> s"
-      using tendsto_unique[of sequentially f l x] trivial_limit_sequentially f(2) by auto
-  }
-  then show "closed s" unfolding closed_limpt by auto
-qed
+proof (unfold closed_sequential_limits, clarify)
+  fix f x assume "\<forall>n. f n \<in> s" and "f ----> x"
+  from `f ----> x` have "Cauchy f"
+    by (rule LIMSEQ_imp_Cauchy)
+  with `complete s` and `\<forall>n. f n \<in> s` obtain l where "l \<in> s" and "f ----> l"
+    by (rule completeE)
+  from `f ----> x` and `f ----> l` have "x = l"
+    by (rule LIMSEQ_unique)
+  with `l \<in> s` show "x \<in> s"
+    by simp
+qed
+
+lemma complete_inter_closed:
+  assumes "complete s" and "closed t"
+  shows "complete (s \<inter> t)"
+proof (rule completeI)
+  fix f assume "\<forall>n. f n \<in> s \<inter> t" and "Cauchy f"
+  then have "\<forall>n. f n \<in> s" and "\<forall>n. f n \<in> t"
+    by simp_all
+  from `complete s` obtain l where "l \<in> s" and "f ----> l"
+    using `\<forall>n. f n \<in> s` and `Cauchy f` by (rule completeE)
+  from `closed t` and `\<forall>n. f n \<in> t` and `f ----> l` have "l \<in> t"
+    by (rule closed_sequentially)
+  with `l \<in> s` and `f ----> l` show "\<exists>l\<in>s \<inter> t. f ----> l"
+    by fast
+qed
+
+lemma complete_closed_subset:
+  assumes "closed s" and "s \<subseteq> t" and "complete t"
+  shows "complete s"
+  using assms complete_inter_closed [of t s] by (simp add: Int_absorb1)
 
 lemma complete_eq_closed:
-  fixes s :: "'a::complete_space set"
-  shows "complete s \<longleftrightarrow> closed s" (is "?lhs = ?rhs")
+  fixes s :: "('a::complete_space) set"
+  shows "complete s \<longleftrightarrow> closed s"
 proof
-  assume ?lhs
-  then show ?rhs by (rule complete_imp_closed)
+  assume "closed s" then show "complete s"
+    using subset_UNIV complete_UNIV by (rule complete_closed_subset)
 next
-  assume ?rhs
-  {
-    fix f
-    assume as:"\<forall>n::nat. f n \<in> s" "Cauchy f"
-    then obtain l where "(f ---> l) sequentially"
-      using complete_univ[unfolded complete_def, THEN spec[where x=f]] by auto
-    then have "\<exists>l\<in>s. (f ---> l) sequentially"
-      using `?rhs`[unfolded closed_sequential_limits, THEN spec[where x=f], THEN spec[where x=l]]
-      using as(1) by auto
-  }
-  then show ?lhs unfolding complete_def by auto
+  assume "complete s" then show "closed s"
+    by (rule complete_imp_closed)
 qed
 
 lemma convergent_eq_cauchy:
@@ -4142,13 +4175,13 @@
 
 lemma compact_cball[simp]:
   fixes x :: "'a::heine_borel"
-  shows "compact(cball x e)"
+  shows "compact (cball x e)"
   using compact_eq_bounded_closed bounded_cball closed_cball
   by blast
 
 lemma compact_frontier_bounded[intro]:
   fixes s :: "'a::heine_borel set"
-  shows "bounded s \<Longrightarrow> compact(frontier s)"
+  shows "bounded s \<Longrightarrow> compact (frontier s)"
   unfolding frontier_def
   using compact_eq_bounded_closed
   by blast
@@ -4168,68 +4201,51 @@
 subsection {* Bounded closed nest property (proof does not use Heine-Borel) *}
 
 lemma bounded_closed_nest:
-  assumes "\<forall>n. closed(s n)"
-    and "\<forall>n. (s n \<noteq> {})"
-    and "(\<forall>m n. m \<le> n --> s n \<subseteq> s m)"
-    and "bounded(s 0)"
-  shows "\<exists>a::'a::heine_borel. \<forall>n::nat. a \<in> s(n)"
+  fixes s :: "nat \<Rightarrow> ('a::heine_borel) set"
+  assumes "\<forall>n. closed (s n)"
+    and "\<forall>n. s n \<noteq> {}"
+    and "\<forall>m n. m \<le> n \<longrightarrow> s n \<subseteq> s m"
+    and "bounded (s 0)"
+  shows "\<exists>a. \<forall>n. a \<in> s n"
 proof -
-  from assms(2) obtain x where x:"\<forall>n::nat. x n \<in> s n"
-    using choice[of "\<lambda>n x. x\<in> s n"] by auto
-  from assms(4,1) have *:"seq_compact (s 0)"
-    using bounded_closed_imp_seq_compact[of "s 0"] by auto
-
-  then obtain l r where lr:"l\<in>s 0" "subseq r" "((x \<circ> r) ---> l) sequentially"
-    unfolding seq_compact_def
-    apply (erule_tac x=x in allE)
-    using x using assms(3)
-    apply blast
-    done
-
-  {
+  from assms(2) obtain x where x: "\<forall>n. x n \<in> s n"
+    using choice[of "\<lambda>n x. x \<in> s n"] by auto
+  from assms(4,1) have "seq_compact (s 0)"
+    by (simp add: bounded_closed_imp_seq_compact)
+  then obtain l r where lr: "l \<in> s 0" "subseq r" "(x \<circ> r) ----> l"
+    using x and assms(3) unfolding seq_compact_def by blast
+  have "\<forall>n. l \<in> s n"
+  proof
     fix n :: nat
-    {
-      fix e :: real
-      assume "e>0"
-      with lr(3) obtain N where N:"\<forall>m\<ge>N. dist ((x \<circ> r) m) l < e"
-        unfolding LIMSEQ_def by auto
-      then have "dist ((x \<circ> r) (max N n)) l < e" by auto
-      moreover
-      have "r (max N n) \<ge> n" using lr(2) using seq_suble[of r "max N n"]
-        by auto
-      then have "(x \<circ> r) (max N n) \<in> s n"
-        using x
-        apply (erule_tac x=n in allE)
-        using x
-        apply (erule_tac x="r (max N n)" in allE)
-        using assms(3)
-        apply (erule_tac x=n in allE)
-        apply (erule_tac x="r (max N n)" in allE)
-        apply auto
-        done
-      ultimately have "\<exists>y\<in>s n. dist y l < e"
-        by auto
-    }
-    then have "l \<in> s n"
-      using closed_approachable[of "s n" l] assms(1) by blast
-  }
-  then show ?thesis by auto
+    have "closed (s n)"
+      using assms(1) by simp
+    moreover have "\<forall>i. (x \<circ> r) i \<in> s i"
+      using x and assms(3) and lr(2) [THEN seq_suble] by auto
+    then have "\<forall>i. (x \<circ> r) (i + n) \<in> s n"
+      using assms(3) by (fast intro!: le_add2)
+    moreover have "(\<lambda>i. (x \<circ> r) (i + n)) ----> l"
+      using lr(3) by (rule LIMSEQ_ignore_initial_segment)
+    ultimately show "l \<in> s n"
+      by (rule closed_sequentially)
+  qed
+  then show ?thesis ..
 qed
 
 text {* Decreasing case does not even need compactness, just completeness. *}
 
 lemma decreasing_closed_nest:
+  fixes s :: "nat \<Rightarrow> ('a::complete_space) set"
   assumes
-    "\<forall>n. closed(s n)"
-    "\<forall>n. (s n \<noteq> {})"
-    "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
-    "\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y \<in> (s n). dist x y < e"
-  shows "\<exists>a::'a::complete_space. \<forall>n::nat. a \<in> s n"
-proof-
-  have "\<forall>n. \<exists> x. x\<in>s n"
+    "\<forall>n. closed (s n)"
+    "\<forall>n. s n \<noteq> {}"
+    "\<forall>m n. m \<le> n \<longrightarrow> s n \<subseteq> s m"
+    "\<forall>e>0. \<exists>n. \<forall>x\<in>s n. \<forall>y\<in>s n. dist x y < e"
+  shows "\<exists>a. \<forall>n. a \<in> s n"
+proof -
+  have "\<forall>n. \<exists>x. x \<in> s n"
     using assms(2) by auto
   then have "\<exists>t. \<forall>n. t n \<in> s n"
-    using choice[of "\<lambda> n x. x \<in> s n"] by auto
+    using choice[of "\<lambda>n x. x \<in> s n"] by auto
   then obtain t where t: "\<forall>n. t n \<in> s n" by auto
   {
     fix e :: real
@@ -4250,7 +4266,7 @@
   then have "Cauchy t"
     unfolding cauchy_def by auto
   then obtain l where l:"(t ---> l) sequentially"
-    using complete_univ unfolding complete_def by auto
+    using complete_UNIV unfolding complete_def by auto
   {
     fix n :: nat
     {
@@ -4285,7 +4301,7 @@
   assumes
     "\<forall>n. closed(s n)"
     "\<forall>n. s n \<noteq> {}"
-    "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
+    "\<forall>m n. m \<le> n \<longrightarrow> s n \<subseteq> s m"
     "\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y\<in>(s n). dist x y < e"
   shows "\<exists>a. \<Inter>(range s) = {a}"
 proof -
@@ -4815,8 +4831,8 @@
   assumes "uniformly_continuous_on s f"
     and "uniformly_continuous_on s g"
   shows "uniformly_continuous_on s (\<lambda>x. f x - g x)"
-  unfolding ab_diff_minus using assms
-  by (intro uniformly_continuous_on_add uniformly_continuous_on_minus)
+  using assms uniformly_continuous_on_add [of s f "- g"]
+    by (simp add: fun_Compl_def uniformly_continuous_on_minus)
 
 text{* Continuity of all kinds is preserved under composition. *}
 
@@ -5637,8 +5653,6 @@
     apply auto
     apply (rule_tac x= xa in exI)
     apply auto
-    apply (rule_tac x=xa in exI)
-    apply auto
     done
   then show ?thesis
     using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto
@@ -5686,24 +5700,27 @@
 
 text {* We can state this in terms of diameter of a set. *}
 
-definition "diameter s = (if s = {} then 0::real else Sup {dist x y | x y. x \<in> s \<and> y \<in> s})"
+definition diameter :: "'a::metric_space set \<Rightarrow> real" where
+  "diameter S = (if S = {} then 0 else SUP (x,y):S\<times>S. dist x y)"
 
 lemma diameter_bounded_bound:
   fixes s :: "'a :: metric_space set"
   assumes s: "bounded s" "x \<in> s" "y \<in> s"
   shows "dist x y \<le> diameter s"
 proof -
-  let ?D = "{dist x y |x y. x \<in> s \<and> y \<in> s}"
   from s obtain z d where z: "\<And>x. x \<in> s \<Longrightarrow> dist z x \<le> d"
     unfolding bounded_def by auto
-  have "dist x y \<le> Sup ?D"
-  proof (rule cSup_upper, safe)
+  have "bdd_above (split dist ` (s\<times>s))"
+  proof (intro bdd_aboveI, safe)
     fix a b
     assume "a \<in> s" "b \<in> s"
     with z[of a] z[of b] dist_triangle[of a b z]
     show "dist a b \<le> 2 * d"
       by (simp add: dist_commute)
-  qed (insert s, auto)
+  qed
+  moreover have "(x,y) \<in> s\<times>s" using s by auto
+  ultimately have "dist x y \<le> (SUP (x,y):s\<times>s. dist x y)"
+    by (rule cSUP_upper2) simp
   with `x \<in> s` show ?thesis
     by (auto simp add: diameter_def)
 qed
@@ -5714,16 +5731,12 @@
     and d: "0 < d" "d < diameter s"
   shows "\<exists>x\<in>s. \<exists>y\<in>s. d < dist x y"
 proof (rule ccontr)
-  let ?D = "{dist x y |x y. x \<in> s \<and> y \<in> s}"
   assume contr: "\<not> ?thesis"
-  moreover
-  from d have "s \<noteq> {}"
-    by (auto simp: diameter_def)
-  then have "?D \<noteq> {}" by auto
-  ultimately have "Sup ?D \<le> d"
-    by (intro cSup_least) (auto simp: not_less)
-  with `d < diameter s` `s \<noteq> {}` show False
-    by (auto simp: diameter_def)
+  moreover have "s \<noteq> {}"
+    using d by (auto simp add: diameter_def)
+  ultimately have "diameter s \<le> d"
+    by (auto simp: not_less diameter_def intro!: cSUP_least)
+  with `d < diameter s` show False by auto
 qed
 
 lemma diameter_bounded:
@@ -5746,7 +5759,7 @@
   then have "diameter s \<le> dist x y"
     unfolding diameter_def
     apply clarsimp
-    apply (rule cSup_least)
+    apply (rule cSUP_least)
     apply fast+
     done
   then show ?thesis
@@ -6989,7 +7002,8 @@
   unfolding homeomorphic_minimal
   apply (rule_tac x="\<lambda>x. a + x" in exI)
   apply (rule_tac x="\<lambda>x. -a + x" in exI)
-  using continuous_on_add[OF continuous_on_const continuous_on_id]
+  using continuous_on_add [OF continuous_on_const continuous_on_id, of s a]
+    continuous_on_add [OF continuous_on_const continuous_on_id, of "plus a ` s" "- a"]
   apply auto
   done
 
@@ -7350,14 +7364,14 @@
     fix y
     assume "a \<le> y" "y \<le> b" "m > 0"
     then have "m *\<^sub>R a + c \<le> m *\<^sub>R y + c" and "m *\<^sub>R y + c \<le> m *\<^sub>R b + c"
-      unfolding eucl_le[where 'a='a] by (auto simp: inner_simps)
+      unfolding eucl_le[where 'a='a] by (auto simp: inner_distrib)
   }
   moreover
   {
     fix y
     assume "a \<le> y" "y \<le> b" "m < 0"
     then have "m *\<^sub>R b + c \<le> m *\<^sub>R y + c" and "m *\<^sub>R y + c \<le> m *\<^sub>R a + c"
-      unfolding eucl_le[where 'a='a] by (auto simp add: mult_left_mono_neg inner_simps)
+      unfolding eucl_le[where 'a='a] by (auto simp add: mult_left_mono_neg inner_distrib)
   }
   moreover
   {
@@ -7366,7 +7380,7 @@
     then have "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
       unfolding image_iff Bex_def mem_interval eucl_le[where 'a='a]
       apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"])
-      apply (auto simp add: pos_le_divide_eq pos_divide_le_eq mult_commute diff_le_iff inner_simps)
+      apply (auto simp add: pos_le_divide_eq pos_divide_le_eq mult_commute diff_le_iff inner_distrib inner_diff_left)
       done
   }
   moreover
@@ -7376,7 +7390,7 @@
     then have "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
       unfolding image_iff Bex_def mem_interval eucl_le[where 'a='a]
       apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"])
-      apply (auto simp add: neg_le_divide_eq neg_divide_le_eq mult_commute diff_le_iff inner_simps)
+      apply (auto simp add: neg_le_divide_eq neg_divide_le_eq mult_commute diff_le_iff inner_distrib inner_diff_left)
       done
   }
   ultimately show ?thesis using False by auto
--- a/src/HOL/NSA/CLim.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/NSA/CLim.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -22,11 +22,11 @@
 lemma all_shift: "(\<forall>x::'a::comm_ring_1. P x) = (\<forall>x. P (x-a))";
 apply auto 
 apply (drule_tac x="x+a" in spec) 
-apply (simp add: diff_minus add_assoc) 
+apply (simp add: add_assoc) 
 done
 
 lemma complex_add_minus_iff [simp]: "(x + - a = (0::complex)) = (x=a)"
-by (simp add: diff_eq_eq diff_minus [symmetric])
+by (simp add: diff_eq_eq)
 
 lemma complex_add_eq_0_iff [iff]: "(x+y = (0::complex)) = (y = -x)"
 apply auto
--- a/src/HOL/NSA/HDeriv.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/NSA/HDeriv.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -81,8 +81,7 @@
 text{*second equivalence *}
 lemma NSDERIV_NSLIM_iff2:
      "(NSDERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --NS> D)"
-by (simp add: NSDERIV_NSLIM_iff DERIV_LIM_iff  diff_minus [symmetric]
-              LIM_NSLIM_iff [symmetric])
+  by (simp add: NSDERIV_NSLIM_iff DERIV_LIM_iff LIM_NSLIM_iff [symmetric])
 
 (* while we're at it! *)
 
@@ -120,11 +119,10 @@
                  hypreal_of_real (f x))\<approx> (hypreal_of_real D) * h)"
 apply (auto simp add: nsderiv_def)
 apply (case_tac "h = (0::hypreal) ")
-apply (auto simp add: diff_minus)
+apply auto
 apply (drule_tac x = h in bspec)
 apply (drule_tac [2] c = h in approx_mult1)
-apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
-            simp add: diff_minus)
+apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD])
 done
 
 lemma NSDERIVD3:
@@ -135,8 +133,7 @@
 apply (auto simp add: nsderiv_def)
 apply (rule ccontr, drule_tac x = h in bspec)
 apply (drule_tac [2] c = h in approx_mult1)
-apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
-            simp add: mult_assoc diff_minus)
+apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: mult_assoc)
 done
 
 text{*Differentiability implies continuity
@@ -174,7 +171,7 @@
 apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def)
 apply (auto simp add: add_divide_distrib diff_divide_distrib dest!: spec)
 apply (drule_tac b = "star_of Da" and d = "star_of Db" in approx_add)
-apply (auto simp add: diff_minus add_ac)
+apply (auto simp add: add_ac algebra_simps)
 done
 
 text{*Product of functions - Proof is trivial but tedious
@@ -234,9 +231,11 @@
   hence deriv: "(\<lambda>h. - ((f(x+h) - f x) / h)) -- 0 --NS> - D"
     by (rule NSLIM_minus)
   have "\<forall>h. - ((f (x + h) - f x) / h) = (- f (x + h) + f x) / h"
-    by (simp add: minus_divide_left diff_minus)
+    by (simp add: minus_divide_left)
   with deriv
-  show "(\<lambda>h. (- f (x + h) + f x) / h) -- 0 --NS> - D" by simp
+  have "(\<lambda>h. (- f (x + h) + f x) / h) -- 0 --NS> - D" by simp
+  then show "(\<lambda>h. (f (x + h) - f x) / h) -- 0 --NS> D \<Longrightarrow>
+    (\<lambda>h. (f x - f (x + h)) / h) -- 0 --NS> - D" by simp
 qed
 
 text{*Subtraction*}
@@ -244,11 +243,8 @@
 by (blast dest: NSDERIV_add NSDERIV_minus)
 
 lemma NSDERIV_diff:
-     "[| NSDERIV f x :> Da; NSDERIV g x :> Db |]
-      ==> NSDERIV (%x. f x - g x) x :> Da-Db"
-apply (simp add: diff_minus)
-apply (blast intro: NSDERIV_add_minus)
-done
+  "NSDERIV f x :> Da \<Longrightarrow> NSDERIV g x :> Db \<Longrightarrow> NSDERIV (\<lambda>x. f x - g x) x :> Da-Db"
+  using NSDERIV_add_minus [of f x Da g Db] by simp
 
 text{*  Similarly to the above, the chain rule admits an entirely
    straightforward derivation. Compare this with Harrison's
@@ -294,7 +290,7 @@
                    - star_of (f (g x)))
               / (( *f* g) (star_of(x) + xa) - star_of (g x))
              \<approx> star_of(Da)"
-by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def diff_minus [symmetric])
+by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def)
 
 (*--------------------------------------------------------------
    from other version of differentiability
@@ -354,13 +350,23 @@
     from h_Inf have "h * star_of x \<in> Infinitesimal" by (rule Infinitesimal_HFinite_mult) simp
     with assms have "inverse (- (h * star_of x) + - (star_of x * star_of x)) \<approx>
       inverse (- (star_of x * star_of x))"
-      by (auto intro: inverse_add_Infinitesimal_approx2
+      apply - apply (rule inverse_add_Infinitesimal_approx2)
+      apply (auto
         dest!: hypreal_of_real_HFinite_diff_Infinitesimal
         simp add: inverse_minus_eq [symmetric] HFinite_minus_iff)
-    with not_0 `h \<noteq> 0` assms have "(inverse (star_of x + h) - inverse (star_of x)) / h \<approx>
+      done
+    moreover from not_0 `h \<noteq> 0` assms
+      have "inverse (- (h * star_of x) + - (star_of x * star_of x)) =
+        (inverse (star_of x + h) - inverse (star_of x)) / h"
+      apply (simp add: division_ring_inverse_diff nonzero_inverse_mult_distrib [symmetric]
+        nonzero_inverse_minus_eq [symmetric] ac_simps ring_distribs)
+      apply (subst nonzero_inverse_minus_eq [symmetric])
+      using distrib_right [symmetric, of h "star_of x" "star_of x"] apply simp
+      apply (simp add: field_simps) 
+      done
+    ultimately have "(inverse (star_of x + h) - inverse (star_of x)) / h \<approx>
       - (inverse (star_of x) * inverse (star_of x))"
-      by (simp add: inverse_add nonzero_inverse_mult_distrib [symmetric]
-        nonzero_inverse_minus_eq [symmetric] add_ac mult_ac diff_minus ring_distribs)
+      using assms by (simp add: nonzero_inverse_mult_distrib [symmetric] nonzero_inverse_minus_eq [symmetric])
   } then show ?thesis by (simp add: nsderiv_def)
 qed
 
--- a/src/HOL/NSA/HLim.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/NSA/HLim.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -71,7 +71,7 @@
 
 lemma NSLIM_diff:
   "\<lbrakk>f -- x --NS> l; g -- x --NS> m\<rbrakk> \<Longrightarrow> (\<lambda>x. f x - g x) -- x --NS> (l - m)"
-by (simp only: diff_minus NSLIM_add NSLIM_minus)
+  by (simp only: NSLIM_add NSLIM_minus diff_conv_add_uminus)
 
 lemma NSLIM_add_minus: "[| f -- x --NS> l; g -- x --NS> m |] ==> (%x. f(x) + -g(x)) -- x --NS> (l + -m)"
 by (simp only: NSLIM_add NSLIM_minus)
@@ -95,7 +95,7 @@
 
 lemma NSLIM_zero_cancel: "(%x. f(x) - l) -- x --NS> 0 ==> f -- x --NS> l"
 apply (drule_tac g = "%x. l" and m = l in NSLIM_add)
-apply (auto simp add: diff_minus add_assoc)
+apply (auto simp add: add_assoc)
 done
 
 lemma NSLIM_const_not_eq:
@@ -243,14 +243,14 @@
 apply (drule_tac [2] x = "- star_of a + x" in spec, safe, simp)
 apply (erule mem_infmal_iff [THEN iffD2, THEN Infinitesimal_add_approx_self [THEN approx_sym]])
 apply (erule_tac [3] approx_minus_iff2 [THEN iffD1])
- prefer 2 apply (simp add: add_commute diff_minus [symmetric])
+ prefer 2 apply (simp add: add_commute)
 apply (rule_tac x = x in star_cases)
 apply (rule_tac [2] x = x in star_cases)
-apply (auto simp add: starfun star_of_def star_n_minus star_n_add add_assoc approx_refl star_n_zero_num)
+apply (auto simp add: starfun star_of_def star_n_minus star_n_add add_assoc star_n_zero_num)
 done
 
 lemma NSLIM_isCont_iff: "(f -- a --NS> f a) = ((%h. f(a + h)) -- 0 --NS> f a)"
-by (rule NSLIM_h_iff)
+  by (fact NSLIM_h_iff)
 
 lemma isNSCont_minus: "isNSCont f a ==> isNSCont (%x. - f x) a"
 by (simp add: isNSCont_def)
--- a/src/HOL/NSA/HSEQ.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/NSA/HSEQ.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -73,14 +73,14 @@
 lemma NSLIMSEQ_minus_cancel: "(%n. -(X n)) ----NS> -a ==> X ----NS> a"
 by (drule NSLIMSEQ_minus, simp)
 
+lemma NSLIMSEQ_diff:
+     "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n - Y n) ----NS> a - b"
+  using NSLIMSEQ_add [of X a "- Y" "- b"] by (simp add: NSLIMSEQ_minus fun_Compl_def)
+
 (* FIXME: delete *)
 lemma NSLIMSEQ_add_minus:
      "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n + -Y n) ----NS> a + -b"
-by (simp add: NSLIMSEQ_add NSLIMSEQ_minus)
-
-lemma NSLIMSEQ_diff:
-     "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n - Y n) ----NS> a - b"
-by (simp add: diff_minus NSLIMSEQ_add NSLIMSEQ_minus)
+  by (simp add: NSLIMSEQ_diff)
 
 lemma NSLIMSEQ_diff_const: "f ----NS> a ==> (%n.(f n - b)) ----NS> a - b"
 by (simp add: NSLIMSEQ_diff NSLIMSEQ_const)
@@ -233,11 +233,11 @@
 
 lemma NSLIMSEQ_inverse_real_of_nat_add_minus:
      "(%n. r + -inverse(real(Suc n))) ----NS> r"
-by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add_minus)
+  using LIMSEQ_inverse_real_of_nat_add_minus by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric])
 
 lemma NSLIMSEQ_inverse_real_of_nat_add_minus_mult:
      "(%n. r*( 1 + -inverse(real(Suc n)))) ----NS> r"
-by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add_minus_mult)
+  using LIMSEQ_inverse_real_of_nat_add_minus_mult by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric])
 
 
 subsection {* Convergence *}
--- a/src/HOL/NSA/HSeries.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/NSA/HSeries.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -131,7 +131,7 @@
 apply (auto simp add: approx_refl)
 apply (drule approx_sym [THEN approx_minus_iff [THEN iffD1]])
 apply (auto dest: approx_hrabs 
-            simp add: sumhr_split_diff diff_minus [symmetric])
+            simp add: sumhr_split_diff)
 done
 
 (*----------------------------------------------------------------
@@ -172,7 +172,7 @@
 apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym])
 apply (rule_tac [2] approx_minus_iff [THEN iffD2])
 apply (auto dest: approx_hrabs_zero_cancel 
-            simp add: sumhr_split_diff diff_minus [symmetric])
+            simp add: sumhr_split_diff)
 done
 
 
--- a/src/HOL/NSA/HTranscendental.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/NSA/HTranscendental.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -258,7 +258,7 @@
             simp add: mult_assoc)
 apply (rule approx_add_right_cancel [where d="-1"])
 apply (rule approx_sym [THEN [2] approx_trans2])
-apply (auto simp add: diff_minus mem_infmal_iff)
+apply (auto simp add: mem_infmal_iff minus_one [symmetric] simp del: minus_one)
 done
 
 lemma STAR_exp_epsilon [simp]: "( *f* exp) epsilon @= 1"
@@ -450,7 +450,7 @@
 apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
             simp add: mult_assoc)
 apply (rule approx_add_right_cancel [where d = "-1"])
-apply (simp add: diff_minus)
+apply (simp add: minus_one [symmetric] del: minus_one)
 done
 
 lemma STAR_tan_zero [simp]: "( *f* tan) 0 = 0"
@@ -587,7 +587,7 @@
      "x \<in> Infinitesimal ==> ( *f* cos) x @= 1 - x\<^sup>2"
 apply (rule STAR_cos_Infinitesimal [THEN approx_trans])
 apply (auto simp add: Infinitesimal_approx_minus [symmetric] 
-            diff_minus add_assoc [symmetric] numeral_2_eq_2)
+            add_assoc [symmetric] numeral_2_eq_2)
 done
 
 lemma STAR_cos_Infinitesimal_approx2:
--- a/src/HOL/NSA/NSA.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/NSA/NSA.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -6,7 +6,7 @@
 header{*Infinite Numbers, Infinitesimals, Infinitely Close Relation*}
 
 theory NSA
-imports HyperDef
+imports HyperDef "~~/src/HOL/Library/Lubs_Glbs"
 begin
 
 definition
@@ -368,7 +368,7 @@
 
 lemma Infinitesimal_diff:
      "[| x \<in> Infinitesimal;  y \<in> Infinitesimal |] ==> x-y \<in> Infinitesimal"
-by (simp add: diff_minus Infinitesimal_add)
+  using Infinitesimal_add [of x "- y"] by simp
 
 lemma Infinitesimal_mult:
   fixes x y :: "'a::real_normed_algebra star"
@@ -491,7 +491,9 @@
      "[|(x::hypreal) \<in> HInfinite; y \<le> 0; x \<le> 0|] ==> (x + y): HInfinite"
 apply (drule HInfinite_minus_iff [THEN iffD2])
 apply (rule HInfinite_minus_iff [THEN iffD1])
-apply (auto intro: HInfinite_add_ge_zero)
+apply (simp only: minus_add add.commute)
+apply (rule HInfinite_add_ge_zero)
+apply simp_all
 done
 
 lemma HInfinite_add_lt_zero:
@@ -620,7 +622,7 @@
 by (simp add: approx_def)
 
 lemma approx_minus_iff2: " (x @= y) = (-y + x @= 0)"
-by (simp add: approx_def diff_minus add_commute)
+by (simp add: approx_def add_commute)
 
 lemma approx_refl [iff]: "x @= x"
 by (simp add: approx_def Infinitesimal_def)
@@ -637,7 +639,7 @@
 lemma approx_trans: "[| x @= y; y @= z |] ==> x @= z"
 apply (simp add: approx_def)
 apply (drule (1) Infinitesimal_add)
-apply (simp add: diff_minus)
+apply simp
 done
 
 lemma approx_trans2: "[| r @= x; s @= x |] ==> r @= s"
@@ -687,7 +689,7 @@
 lemma approx_minus: "a @= b ==> -a @= -b"
 apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym])
 apply (drule approx_minus_iff [THEN iffD1])
-apply (simp add: add_commute diff_minus)
+apply (simp add: add_commute)
 done
 
 lemma approx_minus2: "-a @= -b ==> a @= b"
@@ -700,7 +702,7 @@
 by (blast intro!: approx_add approx_minus)
 
 lemma approx_diff: "[| a @= b; c @= d |] ==> a - c @= b - d"
-by (simp only: diff_minus approx_add approx_minus)
+  using approx_add [of a b "- c" "- d"] by simp
 
 lemma approx_mult1:
   fixes a b c :: "'a::real_normed_algebra star"
@@ -1213,7 +1215,9 @@
          r \<in> Reals;  0 < r |]
       ==> -(x + -t) \<le> r"
 apply (subgoal_tac "(t + -r \<le> x)") 
-apply (auto intro: lemma_st_part_le2)
+apply simp
+apply (rule lemma_st_part_le2)
+apply auto
 done
 
 lemma lemma_SReal_ub:
@@ -1238,7 +1242,7 @@
       ==> x + -t \<noteq> r"
 apply auto
 apply (frule isLubD1a [THEN Reals_minus])
-apply (drule Reals_add_cancel, assumption)
+using Reals_add_cancel [of x "- t"] apply simp
 apply (drule_tac x = x in lemma_SReal_lub)
 apply (drule hypreal_isLub_unique, assumption, auto)
 done
@@ -1250,8 +1254,7 @@
       ==> -(x + -t) \<noteq> r"
 apply (auto)
 apply (frule isLubD1a)
-apply (drule Reals_add_cancel, assumption)
-apply (drule_tac a = "-x" in Reals_minus, simp)
+using Reals_add_cancel [of "- x" t] apply simp
 apply (drule_tac x = x in lemma_SReal_lub)
 apply (drule hypreal_isLub_unique, assumption, auto)
 done
--- a/src/HOL/NSA/NSCA.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/NSA/NSCA.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -165,7 +165,7 @@
 
 lemma approx_hcmod_approx_zero: "(x @= y) = (hcmod (y - x) @= 0)"
 apply (subst hnorm_minus_commute)
-apply (simp add: approx_def Infinitesimal_hcmod_iff diff_minus)
+apply (simp add: approx_def Infinitesimal_hcmod_iff)
 done
 
 lemma approx_approx_zero_iff: "(x @= 0) = (hcmod x @= 0)"
@@ -178,14 +178,14 @@
      "u @= 0 ==> hcmod(x + u) - hcmod x \<in> Infinitesimal"
 apply (drule approx_approx_zero_iff [THEN iffD1])
 apply (rule_tac e = "hcmod u" and e' = "- hcmod u" in Infinitesimal_interval2)
-apply (auto simp add: mem_infmal_iff [symmetric] diff_minus)
+apply (auto simp add: mem_infmal_iff [symmetric])
 apply (rule_tac c1 = "hcmod x" in add_le_cancel_left [THEN iffD1])
-apply (auto simp add: diff_minus [symmetric])
+apply auto
 done
 
 lemma approx_hcmod_add_hcmod: "u @= 0 ==> hcmod(x + u) @= hcmod x"
 apply (rule approx_minus_iff [THEN iffD2])
-apply (auto intro: Infinitesimal_hcmod_add_diff simp add: mem_infmal_iff [symmetric] diff_minus [symmetric])
+apply (auto intro: Infinitesimal_hcmod_add_diff simp add: mem_infmal_iff [symmetric])
 done
 
 
--- a/src/HOL/NSA/StarDef.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/NSA/StarDef.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -803,7 +803,7 @@
 instance star :: (ab_group_add) ab_group_add
 apply (intro_classes)
 apply (transfer, rule left_minus)
-apply (transfer, rule diff_minus)
+apply (transfer, rule diff_conv_add_uminus)
 done
 
 instance star :: (ordered_ab_semigroup_add) ordered_ab_semigroup_add
--- a/src/HOL/Nat.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Nat.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -327,7 +327,7 @@
    apply auto
   done
 
-lemma one_eq_mult_iff [simp,no_atp]: "(Suc 0 = m * n) = (m = Suc 0 & n = Suc 0)"
+lemma one_eq_mult_iff [simp]: "(Suc 0 = m * n) = (m = Suc 0 & n = Suc 0)"
   apply (rule trans)
   apply (rule_tac [2] mult_eq_1_iff, fastforce)
   done
@@ -369,8 +369,8 @@
 | "Suc m \<le> n \<longleftrightarrow> (case n of 0 \<Rightarrow> False | Suc n \<Rightarrow> m \<le> n)"
 
 declare less_eq_nat.simps [simp del]
-lemma [code]: "(0\<Colon>nat) \<le> n \<longleftrightarrow> True" by (simp add: less_eq_nat.simps)
 lemma le0 [iff]: "0 \<le> (n\<Colon>nat)" by (simp add: less_eq_nat.simps)
+lemma [code]: "(0\<Colon>nat) \<le> n \<longleftrightarrow> True" by simp
 
 definition less_nat where
   less_eq_Suc_le: "n < m \<longleftrightarrow> Suc n \<le> m"
@@ -491,7 +491,7 @@
 lemma less_Suc0 [iff]: "(n < Suc 0) = (n = 0)"
   by (simp add: less_Suc_eq)
 
-lemma less_one [iff, no_atp]: "(n < (1::nat)) = (n = 0)"
+lemma less_one [iff]: "(n < (1::nat)) = (n = 0)"
   unfolding One_nat_def by (rule less_Suc0)
 
 lemma Suc_mono: "m < n ==> Suc m < Suc n"
@@ -659,7 +659,7 @@
 lemma gr0_conv_Suc: "(0 < n) = (\<exists>m. n = Suc m)"
 by (fast intro: not0_implies_Suc)
 
-lemma not_gr0 [iff,no_atp]: "!!n::nat. (~ (0 < n)) = (n = 0)"
+lemma not_gr0 [iff]: "!!n::nat. (~ (0 < n)) = (n = 0)"
 using neq0_conv by blast
 
 lemma Suc_le_D: "(Suc n \<le> m') ==> (? m. m' = Suc m)"
@@ -1396,10 +1396,10 @@
 
 text{*Special cases where either operand is zero*}
 
-lemma of_nat_0_eq_iff [simp, no_atp]: "0 = of_nat n \<longleftrightarrow> 0 = n"
+lemma of_nat_0_eq_iff [simp]: "0 = of_nat n \<longleftrightarrow> 0 = n"
   by (fact of_nat_eq_iff [of 0 n, unfolded of_nat_0])
 
-lemma of_nat_eq_0_iff [simp, no_atp]: "of_nat m = 0 \<longleftrightarrow> m = 0"
+lemma of_nat_eq_0_iff [simp]: "of_nat m = 0 \<longleftrightarrow> m = 0"
   by (fact of_nat_eq_iff [of m 0, unfolded of_nat_0])
 
 end
@@ -1432,7 +1432,7 @@
 
 text{*Special cases where either operand is zero*}
 
-lemma of_nat_le_0_iff [simp, no_atp]: "of_nat m \<le> 0 \<longleftrightarrow> m = 0"
+lemma of_nat_le_0_iff [simp]: "of_nat m \<le> 0 \<longleftrightarrow> m = 0"
   by (rule of_nat_le_iff [of _ 0, simplified])
 
 lemma of_nat_0_less_iff [simp]: "0 < of_nat n \<longleftrightarrow> 0 < n"
@@ -1872,12 +1872,12 @@
   shows "m dvd n + q \<longleftrightarrow> m dvd n"
   using assms by (simp add: dvd_plus_eq_right add_commute [of n])
 
-lemma less_dvd_minus:
+lemma less_eq_dvd_minus:
   fixes m n :: nat
-  assumes "m < n"
-  shows "m dvd n \<longleftrightarrow> m dvd (n - m)"
+  assumes "m \<le> n"
+  shows "m dvd n \<longleftrightarrow> m dvd n - m"
 proof -
-  from assms have "n = m + (n - m)" by arith
+  from assms have "n = m + (n - m)" by simp
   then obtain q where "n = m + q" ..
   then show ?thesis by (simp add: dvd_reduce add_commute [of m])
 qed
--- a/src/HOL/Nitpick.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Nitpick.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -33,7 +33,7 @@
 Alternative definitions.
 *}
 
-lemma Ex1_unfold [nitpick_unfold, no_atp]:
+lemma Ex1_unfold [nitpick_unfold]:
 "Ex1 P \<equiv> \<exists>x. {x. P x} = {x}"
 apply (rule eq_reflection)
 apply (simp add: Ex1_def set_eq_iff)
@@ -46,18 +46,18 @@
  apply (erule_tac x = y in allE)
 by auto
 
-lemma rtrancl_unfold [nitpick_unfold, no_atp]: "r\<^sup>* \<equiv> (r\<^sup>+)\<^sup>="
+lemma rtrancl_unfold [nitpick_unfold]: "r\<^sup>* \<equiv> (r\<^sup>+)\<^sup>="
   by (simp only: rtrancl_trancl_reflcl)
 
-lemma rtranclp_unfold [nitpick_unfold, no_atp]:
+lemma rtranclp_unfold [nitpick_unfold]:
 "rtranclp r a b \<equiv> (a = b \<or> tranclp r a b)"
 by (rule eq_reflection) (auto dest: rtranclpD)
 
-lemma tranclp_unfold [nitpick_unfold, no_atp]:
+lemma tranclp_unfold [nitpick_unfold]:
 "tranclp r a b \<equiv> (a, b) \<in> trancl {(x, y). r x y}"
 by (simp add: trancl_def)
 
-lemma [nitpick_simp, no_atp]:
+lemma [nitpick_simp]:
 "of_nat n = (if n = 0 then 0 else 1 + of_nat (n - 1))"
 by (cases n) auto
 
@@ -85,18 +85,18 @@
 \textit{specialize} optimization.
 *}
 
-lemma The_psimp [nitpick_psimp, no_atp]:
+lemma The_psimp [nitpick_psimp]:
   "P = (op =) x \<Longrightarrow> The P = x"
   by auto
 
-lemma Eps_psimp [nitpick_psimp, no_atp]:
+lemma Eps_psimp [nitpick_psimp]:
 "\<lbrakk>P x; \<not> P y; Eps P = y\<rbrakk> \<Longrightarrow> Eps P = x"
 apply (cases "P (Eps P)")
  apply auto
 apply (erule contrapos_np)
 by (rule someI)
 
-lemma unit_case_unfold [nitpick_unfold, no_atp]:
+lemma unit_case_unfold [nitpick_unfold]:
 "unit_case x u \<equiv> x"
 apply (subgoal_tac "u = ()")
  apply (simp only: unit.cases)
@@ -104,14 +104,14 @@
 
 declare unit.cases [nitpick_simp del]
 
-lemma nat_case_unfold [nitpick_unfold, no_atp]:
+lemma nat_case_unfold [nitpick_unfold]:
 "nat_case x f n \<equiv> if n = 0 then x else f (n - 1)"
 apply (rule eq_reflection)
 by (cases n) auto
 
 declare nat.cases [nitpick_simp del]
 
-lemma list_size_simp [nitpick_simp, no_atp]:
+lemma list_size_simp [nitpick_simp]:
 "list_size f xs = (if xs = [] then 0
                    else Suc (f (hd xs) + list_size f (tl xs)))"
 "size xs = (if xs = [] then 0 else Suc (size (tl xs)))"
--- a/src/HOL/Num.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Num.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -407,7 +407,7 @@
   apply (simp add: add_assoc [symmetric], simp add: add_assoc)
   apply (rule_tac a=x in add_left_imp_eq)
   apply (rule_tac a=x in add_right_imp_eq)
-  apply (simp add: add_assoc minus_add_cancel add_minus_cancel)
+  apply (simp add: add_assoc)
   apply (simp add: add_assoc, simp add: add_assoc [symmetric])
   done
 
@@ -418,7 +418,7 @@
 lemmas is_num_normalize =
   add_assoc is_num_add_commute is_num_add_left_commute
   is_num.intros is_num_numeral
-  diff_minus minus_add add_minus_cancel minus_add_cancel
+  minus_add
 
 definition dbl :: "'a \<Rightarrow> 'a" where "dbl x = x + x"
 definition dbl_inc :: "'a \<Rightarrow> 'a" where "dbl_inc x = x + x + 1"
@@ -435,24 +435,21 @@
   "dbl 0 = 0"
   "dbl 1 = 2"
   "dbl (numeral k) = numeral (Bit0 k)"
-  unfolding dbl_def neg_numeral_def numeral.simps
-  by (simp_all add: minus_add)
+  by (simp_all add: dbl_def neg_numeral_def numeral.simps minus_add)
 
 lemma dbl_inc_simps [simp]:
   "dbl_inc (neg_numeral k) = neg_numeral (BitM k)"
   "dbl_inc 0 = 1"
   "dbl_inc 1 = 3"
   "dbl_inc (numeral k) = numeral (Bit1 k)"
-  unfolding dbl_inc_def neg_numeral_def numeral.simps numeral_BitM
-  by (simp_all add: is_num_normalize)
+  by (simp_all add: dbl_inc_def neg_numeral_def numeral.simps numeral_BitM is_num_normalize algebra_simps del: add_uminus_conv_diff)
 
 lemma dbl_dec_simps [simp]:
   "dbl_dec (neg_numeral k) = neg_numeral (Bit1 k)"
   "dbl_dec 0 = -1"
   "dbl_dec 1 = 1"
   "dbl_dec (numeral k) = numeral (BitM k)"
-  unfolding dbl_dec_def neg_numeral_def numeral.simps numeral_BitM
-  by (simp_all add: is_num_normalize)
+  by (simp_all add: dbl_dec_def neg_numeral_def numeral.simps numeral_BitM is_num_normalize)
 
 lemma sub_num_simps [simp]:
   "sub One One = 0"
@@ -464,38 +461,35 @@
   "sub (Bit0 k) (Bit1 l) = dbl_dec (sub k l)"
   "sub (Bit1 k) (Bit0 l) = dbl_inc (sub k l)"
   "sub (Bit1 k) (Bit1 l) = dbl (sub k l)"
-  unfolding dbl_def dbl_dec_def dbl_inc_def sub_def
-  unfolding neg_numeral_def numeral.simps numeral_BitM
-  by (simp_all add: is_num_normalize)
+  by (simp_all add: dbl_def dbl_dec_def dbl_inc_def sub_def neg_numeral_def numeral.simps
+    numeral_BitM is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus)
 
 lemma add_neg_numeral_simps:
   "numeral m + neg_numeral n = sub m n"
   "neg_numeral m + numeral n = sub n m"
   "neg_numeral m + neg_numeral n = neg_numeral (m + n)"
-  unfolding sub_def diff_minus neg_numeral_def numeral_add numeral.simps
-  by (simp_all add: is_num_normalize)
+  by (simp_all add: sub_def neg_numeral_def numeral_add numeral.simps is_num_normalize
+    del: add_uminus_conv_diff add: diff_conv_add_uminus)
 
 lemma add_neg_numeral_special:
   "1 + neg_numeral m = sub One m"
   "neg_numeral m + 1 = sub One m"
-  unfolding sub_def diff_minus neg_numeral_def numeral_add numeral.simps
-  by (simp_all add: is_num_normalize)
+  by (simp_all add: sub_def neg_numeral_def numeral_add numeral.simps is_num_normalize)
 
 lemma diff_numeral_simps:
   "numeral m - numeral n = sub m n"
   "numeral m - neg_numeral n = numeral (m + n)"
   "neg_numeral m - numeral n = neg_numeral (m + n)"
   "neg_numeral m - neg_numeral n = sub n m"
-  unfolding neg_numeral_def sub_def diff_minus numeral_add numeral.simps
-  by (simp_all add: is_num_normalize)
+  by (simp_all add: neg_numeral_def sub_def numeral_add numeral.simps is_num_normalize
+    del: add_uminus_conv_diff add: diff_conv_add_uminus)
 
 lemma diff_numeral_special:
   "1 - numeral n = sub One n"
   "1 - neg_numeral n = numeral (One + n)"
   "numeral m - 1 = sub m One"
   "neg_numeral m - 1 = neg_numeral (m + One)"
-  unfolding neg_numeral_def sub_def diff_minus numeral_add numeral.simps
-  by (simp_all add: is_num_normalize)
+  by (simp_all add: neg_numeral_def sub_def numeral_add numeral.simps add: is_num_normalize)
 
 lemma minus_one: "- 1 = -1"
   unfolding neg_numeral_def numeral.simps ..
@@ -1078,6 +1072,16 @@
   BitM.simps dbl_simps dbl_inc_simps dbl_dec_simps
   abs_zero abs_one arith_extra_simps
 
+lemmas more_arith_simps =
+  neg_le_iff_le
+  minus_zero left_minus right_minus
+  mult_1_left mult_1_right
+  mult_minus_left mult_minus_right
+  minus_add_distrib minus_minus mult_assoc
+
+lemmas of_nat_simps =
+  of_nat_0 of_nat_1 of_nat_Suc of_nat_add of_nat_mult
+
 text {* Simplification of relational operations *}
 
 lemmas eq_numeral_extra =
@@ -1089,6 +1093,38 @@
   less_numeral_simps less_neg_numeral_simps less_numeral_extra
   eq_numeral_simps eq_neg_numeral_simps eq_numeral_extra
 
+lemma Let_numeral [simp]: "Let (numeral v) f = f (numeral v)"
+  -- {* Unfold all @{text let}s involving constants *}
+  unfolding Let_def ..
+
+lemma Let_neg_numeral [simp]: "Let (neg_numeral v) f = f (neg_numeral v)"
+  -- {* Unfold all @{text let}s involving constants *}
+  unfolding Let_def ..
+
+declaration {*
+let 
+  fun number_of thy T n =
+    if not (Sign.of_sort thy (T, @{sort numeral}))
+    then raise CTERM ("number_of", [])
+    else Numeral.mk_cnumber (Thm.ctyp_of thy T) n;
+in
+  K (
+    Lin_Arith.add_simps (@{thms arith_simps} @ @{thms more_arith_simps}
+      @ @{thms rel_simps}
+      @ @{thms pred_numeral_simps}
+      @ @{thms arith_special numeral_One}
+      @ @{thms of_nat_simps})
+    #> Lin_Arith.add_simps [@{thm Suc_numeral},
+      @{thm Let_numeral}, @{thm Let_neg_numeral}, @{thm Let_0}, @{thm Let_1},
+      @{thm le_Suc_numeral}, @{thm le_numeral_Suc},
+      @{thm less_Suc_numeral}, @{thm less_numeral_Suc},
+      @{thm Suc_eq_numeral}, @{thm eq_numeral_Suc},
+      @{thm mult_Suc}, @{thm mult_Suc_right},
+      @{thm of_nat_numeral}]
+    #> Lin_Arith.set_number_of number_of)
+end
+*}
+
 
 subsubsection {* Simplification of arithmetic when nested to the right. *}
 
@@ -1119,4 +1155,3 @@
 
 end
 
-
--- a/src/HOL/Number_Theory/Cong.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Number_Theory/Cong.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -543,7 +543,8 @@
   apply (subgoal_tac "a * b = (-a * -b)")
   apply (erule ssubst)
   apply (subst zmod_zmult2_eq)
-  apply (auto simp add: mod_add_left_eq)
+  apply (auto simp add: mod_add_left_eq mod_minus_right div_minus_right)
+  apply (metis mod_diff_left_eq mod_diff_right_eq mod_mult_self1_is_0 semiring_numeral_div_class.diff_zero)+
   done
 
 lemma cong_to_1_nat: "([(a::nat) = 1] (mod n)) \<Longrightarrow> (n dvd (a - 1))"
--- a/src/HOL/Number_Theory/Eratosthenes.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Number_Theory/Eratosthenes.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -75,7 +75,7 @@
 lemma numbers_of_marks_mark_out:
   "numbers_of_marks n (mark_out m bs) = {q \<in> numbers_of_marks n bs. \<not> Suc m dvd Suc q - n}"
   by (auto simp add: numbers_of_marks_def mark_out_def in_set_enumerate_eq image_iff
-    nth_enumerate_eq less_dvd_minus)
+    nth_enumerate_eq less_eq_dvd_minus)
 
 
 text {* Auxiliary operation for efficient implementation  *}
@@ -128,7 +128,7 @@
     by (simp add: mark_out_aux_def)
   show ?thesis2
     by (auto simp add: mark_out_code [symmetric] mark_out_aux_def mark_out_def
-      enumerate_Suc_eq in_set_enumerate_eq less_dvd_minus)
+      enumerate_Suc_eq in_set_enumerate_eq less_eq_dvd_minus)
   { def v \<equiv> "Suc m" and w \<equiv> "Suc n"
     fix q
     assume "m + n \<le> q"
--- a/src/HOL/Number_Theory/Primes.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Number_Theory/Primes.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -74,8 +74,9 @@
 subsection {* Primes *}
 
 lemma prime_odd_nat: "prime (p::nat) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
-  unfolding prime_nat_def
-  by (metis gcd_lcm_complete_lattice_nat.bot_least nat_even_iff_2_dvd nat_neq_iff odd_1_nat)
+  apply (auto simp add: prime_nat_def even_def dvd_eq_mod_eq_0)
+  apply (metis dvd_eq_mod_eq_0 even_Suc even_def mod_by_1 nat_dvd_not_less not_mod_2_eq_0_eq_1 zero_less_numeral)
+  done
 
 lemma prime_odd_int: "prime (p::int) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
   unfolding prime_int_def
--- a/src/HOL/Numeral_Simprocs.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Numeral_Simprocs.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -13,7 +13,7 @@
 ML_file "~~/src/Provers/Arith/extract_common_term.ML"
 
 lemmas semiring_norm =
-  Let_def arith_simps nat_arith rel_simps
+  Let_def arith_simps diff_nat_numeral rel_simps
   if_False if_True
   add_0 add_Suc add_numeral_left
   add_neg_numeral_left mult_numeral_left
@@ -278,27 +278,8 @@
   ("((l::nat) * m) dvd n" | "(l::nat) dvd (m * n)") =
   {* fn phi => Nat_Numeral_Simprocs.dvd_cancel_factor *}
 
-(* FIXME: duplicate rule warnings for:
-  ring_distribs
-  numeral_plus_numeral numeral_times_numeral
-  numeral_eq_iff numeral_less_iff numeral_le_iff
-  numeral_neq_zero zero_neq_numeral zero_less_numeral
-  if_True if_False *)
 declaration {* 
-  K (Lin_Arith.add_simps ([@{thm Suc_numeral}, @{thm int_numeral}])
-  #> Lin_Arith.add_simps (@{thms ring_distribs} @ [@{thm Let_numeral}, @{thm Let_neg_numeral}, @{thm Let_0}, @{thm Let_1},
-     @{thm nat_0}, @{thm nat_1},
-     @{thm numeral_plus_numeral}, @{thm diff_nat_numeral}, @{thm numeral_times_numeral},
-     @{thm numeral_eq_iff}, @{thm numeral_less_iff}, @{thm numeral_le_iff},
-     @{thm le_Suc_numeral}, @{thm le_numeral_Suc},
-     @{thm less_Suc_numeral}, @{thm less_numeral_Suc},
-     @{thm Suc_eq_numeral}, @{thm eq_numeral_Suc},
-     @{thm mult_Suc}, @{thm mult_Suc_right},
-     @{thm add_Suc}, @{thm add_Suc_right},
-     @{thm numeral_neq_zero}, @{thm zero_neq_numeral}, @{thm zero_less_numeral},
-     @{thm of_int_numeral}, @{thm of_nat_numeral}, @{thm nat_numeral},
-     @{thm if_True}, @{thm if_False}])
-  #> Lin_Arith.add_simprocs
+  K (Lin_Arith.add_simprocs
       [@{simproc semiring_assoc_fold},
        @{simproc int_combine_numerals},
        @{simproc inteq_cancel_numerals},
--- a/src/HOL/Old_Number_Theory/Pocklington.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Old_Number_Theory/Pocklington.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -348,15 +348,11 @@
   let ?w = "x mod (a*b)"
   have wab: "?w < a*b" by (simp add: mod_less_divisor[OF abpos])
   from xq12(1) have "?w mod a = ((m + q1 * a) mod (a*b)) mod a" by simp
-  also have "\<dots> = m mod a" apply (simp add: mod_mult2_eq)
-    apply (subst mod_add_left_eq)
-    by simp
+  also have "\<dots> = m mod a" by (simp add: mod_mult2_eq)
   finally have th1: "[?w = m] (mod a)" by (simp add: modeq_def)
   from xq12(2) have "?w mod b = ((n + q2 * b) mod (a*b)) mod b" by simp
   also have "\<dots> = ((n + q2 * b) mod (b*a)) mod b" by (simp add: mult_commute)
-  also have "\<dots> = n mod b" apply (simp add: mod_mult2_eq)
-    apply (subst mod_add_left_eq)
-    by simp
+  also have "\<dots> = n mod b" by (simp add: mod_mult2_eq)
   finally have th2: "[?w = n] (mod b)" by (simp add: modeq_def)
   {fix y assume H: "y < a*b" "[y = m] (mod a)" "[y = n] (mod b)"
     with th1 th2 have H': "[y = ?w] (mod a)" "[y = ?w] (mod b)"
--- a/src/HOL/Orderings.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Orderings.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -998,7 +998,7 @@
     (!!x y. x > y ==> f x > f y) ==> f a > c"
 by (subgoal_tac "f a > f b", force, force)
 
-lemmas xtrans = xt1 xt2 xt3 xt4 xt5 xt6 xt7 xt8 xt9 [no_atp]
+lemmas xtrans = xt1 xt2 xt3 xt4 xt5 xt6 xt7 xt8 xt9
 
 (* 
   Since "a >= b" abbreviates "b <= a", the abbreviation "..." stands
--- a/src/HOL/Parity.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Parity.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -9,26 +9,52 @@
 imports Main
 begin
 
-class even_odd = 
-  fixes even :: "'a \<Rightarrow> bool"
-
-abbreviation
-  odd :: "'a\<Colon>even_odd \<Rightarrow> bool" where
-  "odd x \<equiv> \<not> even x"
-
-instantiation nat and int  :: even_odd
+class even_odd = semiring_div_parity
 begin
 
-definition
-  even_def [presburger]: "even x \<longleftrightarrow> (x\<Colon>int) mod 2 = 0"
+definition even :: "'a \<Rightarrow> bool"
+where
+  even_def [presburger]: "even a \<longleftrightarrow> a mod 2 = 0"
+
+lemma even_iff_2_dvd [algebra]:
+  "even a \<longleftrightarrow> 2 dvd a"
+  by (simp add: even_def dvd_eq_mod_eq_0)
+
+lemma even_zero [simp]:
+  "even 0"
+  by (simp add: even_def)
+
+lemma even_times_anything:
+  "even a \<Longrightarrow> even (a * b)"
+  by (simp add: even_iff_2_dvd)
 
-definition
-  even_nat_def [presburger]: "even x \<longleftrightarrow> even (int x)"
+lemma anything_times_even:
+  "even a \<Longrightarrow> even (b * a)"
+  by (simp add: even_iff_2_dvd)
+
+abbreviation odd :: "'a \<Rightarrow> bool"
+where
+  "odd a \<equiv> \<not> even a"
 
-instance ..
+lemma odd_times_odd:
+  "odd a \<Longrightarrow> odd b \<Longrightarrow> odd (a * b)" 
+  by (auto simp add: even_def mod_mult_left_eq)
+
+lemma even_product [simp, presburger]:
+  "even (a * b) \<longleftrightarrow> even a \<or> even b"
+  apply (auto simp add: even_times_anything anything_times_even)
+  apply (rule ccontr)
+  apply (auto simp add: odd_times_odd)
+  done
 
 end
 
+instance nat and int  :: even_odd ..
+
+lemma even_nat_def [presburger]:
+  "even x \<longleftrightarrow> even (int x)"
+  by (auto simp add: even_def int_eq_iff int_mult nat_mult_distrib)
+  
 lemma transfer_int_nat_relations:
   "even (int x) \<longleftrightarrow> even x"
   by (simp add: even_nat_def)
@@ -37,13 +63,13 @@
   transfer_int_nat_relations
 ]
 
-lemma even_zero_int[simp]: "even (0::int)" by presburger
-
-lemma odd_one_int[simp]: "odd (1::int)" by presburger
+lemma odd_one_int [simp]:
+  "odd (1::int)"
+  by presburger
 
-lemma even_zero_nat[simp]: "even (0::nat)" by presburger
-
-lemma odd_1_nat [simp]: "odd (1::nat)" by presburger
+lemma odd_1_nat [simp]:
+  "odd (1::nat)"
+  by presburger
 
 lemma even_numeral_int [simp]: "even (numeral (Num.Bit0 k) :: int)"
   unfolding even_def by simp
@@ -52,7 +78,7 @@
   unfolding even_def by simp
 
 (* TODO: proper simp rules for Num.Bit0, Num.Bit1 *)
-declare even_def[of "neg_numeral v", simp] for v
+declare even_def [of "neg_numeral v", simp] for v
 
 lemma even_numeral_nat [simp]: "even (numeral (Num.Bit0 k) :: nat)"
   unfolding even_nat_def by simp
@@ -62,34 +88,8 @@
 
 subsection {* Even and odd are mutually exclusive *}
 
-lemma int_pos_lt_two_imp_zero_or_one:
-    "0 <= x ==> (x::int) < 2 ==> x = 0 | x = 1"
-  by presburger
-
-lemma neq_one_mod_two [simp, presburger]: 
-  "((x::int) mod 2 ~= 0) = (x mod 2 = 1)" by presburger
-
 
 subsection {* Behavior under integer arithmetic operations *}
-declare dvd_def[algebra]
-lemma nat_even_iff_2_dvd[algebra]: "even (x::nat) \<longleftrightarrow> 2 dvd x"
-  by presburger
-lemma int_even_iff_2_dvd[algebra]: "even (x::int) \<longleftrightarrow> 2 dvd x"
-  by presburger
-
-lemma even_times_anything: "even (x::int) ==> even (x * y)"
-  by algebra
-
-lemma anything_times_even: "even (y::int) ==> even (x * y)" by algebra
-
-lemma odd_times_odd: "odd (x::int) ==> odd y ==> odd (x * y)" 
-  by (simp add: even_def mod_mult_right_eq)
-
-lemma even_product[simp,presburger]: "even((x::int) * y) = (even x | even y)"
-  apply (auto simp add: even_times_anything anything_times_even)
-  apply (rule ccontr)
-  apply (auto simp add: odd_times_odd)
-  done
 
 lemma even_plus_even: "even (x::int) ==> even y ==> even (x + y)"
 by presburger
@@ -158,10 +158,6 @@
 
 subsection {* Equivalent definitions *}
 
-lemma nat_lt_two_imp_zero_or_one:
-  "(x::nat) < Suc (Suc 0) ==> x = 0 | x = Suc 0"
-by presburger
-
 lemma even_nat_mod_two_eq_zero: "even (x::nat) ==> x mod (Suc (Suc 0)) = 0"
 by presburger
 
@@ -189,45 +185,19 @@
 
 subsection {* Parity and powers *}
 
-lemma  minus_one_even_odd_power:
-     "(even x --> (- 1::'a::{comm_ring_1})^x = 1) &
-      (odd x --> (- 1::'a)^x = - 1)"
-  apply (induct x)
-  apply (rule conjI)
-  apply simp
-  apply (insert even_zero_nat, blast)
-  apply simp
-  done
-
-lemma minus_one_even_power [simp]:
-    "even x ==> (- 1::'a::{comm_ring_1})^x = 1"
-  using minus_one_even_odd_power by blast
-
-lemma minus_one_odd_power [simp]:
-    "odd x ==> (- 1::'a::{comm_ring_1})^x = - 1"
-  using minus_one_even_odd_power by blast
+lemma (in comm_ring_1) neg_power_if:
+  "(- a) ^ n = (if even n then (a ^ n) else - (a ^ n))"
+  by (induct n) simp_all
 
-lemma neg_one_even_odd_power:
-     "(even x --> (-1::'a::{comm_ring_1})^x = 1) &
-      (odd x --> (-1::'a)^x = -1)"
-  apply (induct x)
-  apply (simp, simp)
-  done
-
-lemma neg_one_even_power [simp]:
-    "even x ==> (-1::'a::{comm_ring_1})^x = 1"
-  using neg_one_even_odd_power by blast
+lemma (in comm_ring_1)
+  shows minus_one_even_power [simp]: "even n \<Longrightarrow> (- 1) ^ n = 1"
+  and minus_one_odd_power [simp]: "odd n \<Longrightarrow> (- 1) ^ n = - 1"
+  by (simp_all add: neg_power_if del: minus_one)
 
-lemma neg_one_odd_power [simp]:
-    "odd x ==> (-1::'a::{comm_ring_1})^x = -1"
-  using neg_one_even_odd_power by blast
-
-lemma neg_power_if:
-     "(-x::'a::{comm_ring_1}) ^ n =
-      (if even n then (x ^ n) else -(x ^ n))"
-  apply (induct n)
-  apply simp_all
-  done
+lemma (in comm_ring_1)
+  shows neg_one_even_power [simp]: "even n \<Longrightarrow> (-1) ^ n = 1"
+  and neg_one_odd_power [simp]: "odd n \<Longrightarrow> (-1) ^ n = - 1"
+  by (simp_all add: minus_one [symmetric] del: minus_one)
 
 lemma zero_le_even_power: "even n ==>
     0 <= (x::'a::{linordered_ring,monoid_mult}) ^ n"
@@ -244,7 +214,7 @@
 apply (metis field_power_not_zero divisors_zero order_antisym_conv zero_le_square)
 done
 
-lemma zero_le_power_eq[presburger]: "(0 <= (x::'a::{linordered_idom}) ^ n) =
+lemma zero_le_power_eq [presburger]: "(0 <= (x::'a::{linordered_idom}) ^ n) =
     (even n | (odd n & 0 <= x))"
   apply auto
   apply (subst zero_le_odd_power [symmetric])
@@ -277,9 +247,6 @@
   apply (simp add: zero_le_even_power)
   done
 
-lemma zero_less_power_nat_eq[presburger]: "(0 < (x::nat) ^ n) = (n = 0 | 0 < x)"
-  by (induct n) auto
-
 lemma power_minus_even [simp]: "even n ==>
     (- x)^n = (x^n::'a::{comm_ring_1})"
   apply (subst power_minus)
@@ -336,13 +303,6 @@
 
 lemma odd_add [simp]: "odd(m + n::nat) = (odd m \<noteq> odd n)" by presburger
 
-lemma div_Suc: "Suc a div c = a div c + Suc 0 div c +
-    (a mod c + Suc 0 mod c) div c" 
-  apply (subgoal_tac "Suc a = a + Suc 0")
-  apply (erule ssubst)
-  apply (rule div_add1_eq, simp)
-  done
-
 lemma lemma_even_div2 [simp]: "even (n::nat) ==> (n + 1) div 2 = n div 2" by presburger
 
 lemma lemma_not_even_div2 [simp]: "~even n ==> (n + 1) div 2 = Suc (n div 2)"
@@ -359,31 +319,29 @@
 text {* Simplify, when the exponent is a numeral *}
 
 lemmas zero_le_power_eq_numeral [simp] =
-    zero_le_power_eq [of _ "numeral w"] for w
+  zero_le_power_eq [of _ "numeral w"] for w
 
 lemmas zero_less_power_eq_numeral [simp] =
-    zero_less_power_eq [of _ "numeral w"] for w
+  zero_less_power_eq [of _ "numeral w"] for w
 
 lemmas power_le_zero_eq_numeral [simp] =
-    power_le_zero_eq [of _ "numeral w"] for w
+  power_le_zero_eq [of _ "numeral w"] for w
 
 lemmas power_less_zero_eq_numeral [simp] =
-    power_less_zero_eq [of _ "numeral w"] for w
+  power_less_zero_eq [of _ "numeral w"] for w
 
 lemmas zero_less_power_nat_eq_numeral [simp] =
-    zero_less_power_nat_eq [of _ "numeral w"] for w
+  nat_zero_less_power_iff [of _ "numeral w"] for w
 
-lemmas power_eq_0_iff_numeral [simp] = power_eq_0_iff [of _ "numeral w"] for w
+lemmas power_eq_0_iff_numeral [simp] =
+  power_eq_0_iff [of _ "numeral w"] for w
 
-lemmas power_even_abs_numeral [simp] = power_even_abs [of "numeral w" _] for w
+lemmas power_even_abs_numeral [simp] =
+  power_even_abs [of "numeral w" _] for w
 
 
 subsection {* An Equivalence for @{term [source] "0 \<le> a^n"} *}
 
-lemma even_power_le_0_imp_0:
-    "a ^ (2*k) \<le> (0::'a::{linordered_idom}) ==> a=0"
-  by (induct k) (auto simp add: zero_le_mult_iff mult_le_0_iff)
-
 lemma zero_le_power_iff[presburger]:
   "(0 \<le> a^n) = (0 \<le> (a::'a::{linordered_idom}) | even n)"
 proof cases
@@ -395,9 +353,10 @@
   assume odd: "odd n"
   then obtain k where "n = Suc(2*k)"
     by (auto simp add: odd_nat_equiv_def2 numeral_2_eq_2)
-  thus ?thesis
-    by (auto simp add: zero_le_mult_iff zero_le_even_power
-             dest!: even_power_le_0_imp_0)
+  moreover have "a ^ (2 * k) \<le> 0 \<Longrightarrow> a = 0"
+    by (induct k) (auto simp add: zero_le_mult_iff mult_le_0_iff)
+  ultimately show ?thesis
+    by (auto simp add: zero_le_mult_iff zero_le_even_power)
 qed
 
 
@@ -409,7 +368,6 @@
 lemma odd_plus_one_div_two: "odd (x::int) ==> (x + 1) div 2 = x div 2 + 1" by presburger
 
 lemma [presburger]: "(Suc x) div Suc (Suc 0) = x div Suc (Suc 0) \<longleftrightarrow> even x" by presburger
-lemma [presburger]: "(Suc x) div Suc (Suc 0) = x div Suc (Suc 0) \<longleftrightarrow> even x" by presburger
 lemma even_nat_plus_one_div_two: "even (x::nat) ==>
     (Suc x) div Suc (Suc 0) = x div Suc (Suc 0)" by presburger
 
@@ -417,3 +375,4 @@
     (Suc x) div Suc (Suc 0) = Suc (x div Suc (Suc 0))" by presburger
 
 end
+
--- a/src/HOL/Power.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Power.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -530,7 +530,7 @@
   "abs ((-a) ^ n) = abs (a ^ n)"
   by (simp add: power_abs)
 
-lemma zero_less_power_abs_iff [simp, no_atp]:
+lemma zero_less_power_abs_iff [simp]:
   "0 < abs a ^ n \<longleftrightarrow> a \<noteq> 0 \<or> n = 0"
 proof (induct n)
   case 0 show ?case by simp
@@ -730,8 +730,18 @@
   fixes m n :: nat
   assumes "m\<^sup>2 \<le> n"
   shows "m \<le> n"
-  using assms by (cases m) (simp_all add: power2_eq_square)
-
+proof (cases m)
+  case 0 then show ?thesis by simp
+next
+  case (Suc k)
+  show ?thesis
+  proof (rule ccontr)
+    assume "\<not> m \<le> n"
+    then have "n < m" by simp
+    with assms Suc show False
+      by (auto simp add: algebra_simps) (simp add: power2_eq_square)
+  qed
+qed
 
 
 subsection {* Code generator tweak *}
--- a/src/HOL/Presburger.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Presburger.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -360,11 +360,15 @@
   apply simp
   done
 
-lemma zdvd_mono: assumes not0: "(k::int) \<noteq> 0"
-shows "((m::int) dvd t) \<equiv> (k*m dvd k*t)" 
-  using not0 by (simp add: dvd_def)
+lemma zdvd_mono:
+  fixes k m t :: int
+  assumes "k \<noteq> 0"
+  shows "m dvd t \<equiv> k * m dvd k * t" 
+  using assms by simp
 
-lemma uminus_dvd_conv: "(d dvd (t::int)) \<equiv> (-d dvd t)" "(d dvd (t::int)) \<equiv> (d dvd -t)"
+lemma uminus_dvd_conv:
+  fixes d t :: int
+  shows "d dvd t \<equiv> - d dvd t" and "d dvd t \<equiv> d dvd - t"
   by simp_all
 
 text {* \bigskip Theorems for transforming predicates on nat to predicates on @{text int}*}
@@ -406,24 +410,23 @@
   end
 *} "Cooper's algorithm for Presburger arithmetic"
 
-declare dvd_eq_mod_eq_0[symmetric, presburger]
-declare mod_1[presburger] 
-declare mod_0[presburger]
-declare mod_by_1[presburger]
-declare mod_self[presburger]
-declare mod_by_0[presburger]
-declare mod_div_trivial[presburger]
-declare div_mod_equality2[presburger]
-declare div_mod_equality[presburger]
-declare mod_div_equality2[presburger]
-declare mod_div_equality[presburger]
-declare mod_mult_self1[presburger]
-declare mod_mult_self2[presburger]
-declare div_mod_equality[presburger]
-declare div_mod_equality2[presburger]
+declare dvd_eq_mod_eq_0 [symmetric, presburger]
+declare mod_1 [presburger] 
+declare mod_0 [presburger]
+declare mod_by_1 [presburger]
+declare mod_self [presburger]
+declare div_by_0 [presburger]
+declare mod_by_0 [presburger]
+declare mod_div_trivial [presburger]
+declare div_mod_equality2 [presburger]
+declare div_mod_equality [presburger]
+declare mod_div_equality2 [presburger]
+declare mod_div_equality [presburger]
+declare mod_mult_self1 [presburger]
+declare mod_mult_self2 [presburger]
 declare mod2_Suc_Suc[presburger]
-lemma [presburger]: "(a::int) div 0 = 0" and [presburger]: "a mod 0 = a"
-by simp_all
+declare not_mod_2_eq_0_eq_1 [presburger] 
+declare nat_zero_less_power_iff [presburger]
 
 lemma [presburger, algebra]: "m mod 2 = (1::nat) \<longleftrightarrow> \<not> 2 dvd m " by presburger
 lemma [presburger, algebra]: "m mod 2 = Suc 0 \<longleftrightarrow> \<not> 2 dvd m " by presburger
@@ -432,3 +435,4 @@
 lemma [presburger, algebra]: "m mod 2 = (1::int) \<longleftrightarrow> \<not> 2 dvd m " by presburger
 
 end
+
--- a/src/HOL/Probability/Borel_Space.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Probability/Borel_Space.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -677,7 +677,7 @@
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
-  unfolding diff_minus using assms by simp
+  using borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def)
 
 lemma borel_measurable_times[measurable (raw)]:
   fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_algebra}"
@@ -719,7 +719,8 @@
   proof cases
     assume "b \<noteq> 0"
     with `open S` have "open ((\<lambda>x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S")
-      by (auto intro!: open_affinity simp: scaleR_add_right)
+      using open_affinity [of S "inverse b" "- a /\<^sub>R b"]
+      by (auto simp: algebra_simps)
     hence "?S \<in> sets borel" by auto
     moreover
     from `b \<noteq> 0` have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S"
--- a/src/HOL/Probability/Lebesgue_Integration.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Probability/Lebesgue_Integration.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1528,7 +1528,7 @@
     using mono by auto
   ultimately show ?thesis using fg
     by (auto intro!: add_mono positive_integral_mono_AE real_of_ereal_positive_mono
-             simp: positive_integral_positive lebesgue_integral_def diff_minus)
+             simp: positive_integral_positive lebesgue_integral_def algebra_simps)
 qed
 
 lemma integral_mono:
@@ -1732,7 +1732,7 @@
   shows "integrable M (\<lambda>t. f t - g t)"
   and "(\<integral> t. f t - g t \<partial>M) = integral\<^sup>L M f - integral\<^sup>L M g"
   using integral_add[OF f integral_minus(1)[OF g]]
-  unfolding diff_minus integral_minus(2)[OF g]
+  unfolding integral_minus(2)[OF g]
   by auto
 
 lemma integral_indicator[simp, intro]:
--- a/src/HOL/Probability/Lebesgue_Measure.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Probability/Lebesgue_Measure.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1138,7 +1138,7 @@
     show "\<And>i x. 0 \<le> ?f i x"
       using nonneg by (auto split: split_indicator)
   qed
-  also have "\<dots> = (SUP i::nat. F (a + real i) - F a)"
+  also have "\<dots> = (SUP i::nat. ereal (F (a + real i) - F a))"
     by (subst positive_integral_FTC_atLeastAtMost[OF f_borel f nonneg]) auto
   also have "\<dots> = T - F a"
   proof (rule SUP_Lim_ereal)
--- a/src/HOL/Product_Type.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Product_Type.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -75,10 +75,10 @@
   f} rather than by @{term [source] "%u. f ()"}.
 *}
 
-lemma unit_abs_eta_conv [simp, no_atp]: "(%u::unit. f ()) = f"
+lemma unit_abs_eta_conv [simp]: "(%u::unit. f ()) = f"
   by (rule ext) simp
 
-lemma UNIV_unit [no_atp]:
+lemma UNIV_unit:
   "UNIV = {()}" by auto
 
 instantiation unit :: default
@@ -586,10 +586,10 @@
    to quite time-consuming computations (in particular for nested tuples) *)
 setup {* map_theory_claset (fn ctxt => ctxt addSbefore ("split_conv_tac", split_conv_tac)) *}
 
-lemma split_eta_SetCompr [simp,no_atp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P"
+lemma split_eta_SetCompr [simp, no_atp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P"
   by (rule ext) fast
 
-lemma split_eta_SetCompr2 [simp,no_atp]: "(%u. EX x y. u = (x, y) & P x y) = split P"
+lemma split_eta_SetCompr2 [simp, no_atp]: "(%u. EX x y. u = (x, y) & P x y) = split P"
   by (rule ext) fast
 
 lemma split_part [simp]: "(%(a,b). P & Q a b) = (%ab. P & split Q ab)"
--- a/src/HOL/Proofs/Lambda/ListOrder.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Proofs/Lambda/ListOrder.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -89,15 +89,15 @@
   done
 
 lemma Cons_acc_step1I [intro!]:
-    "accp r x ==> accp (step1 r) xs \<Longrightarrow> accp (step1 r) (x # xs)"
-  apply (induct arbitrary: xs set: accp)
+    "Wellfounded.accp r x ==> Wellfounded.accp (step1 r) xs \<Longrightarrow> Wellfounded.accp (step1 r) (x # xs)"
+  apply (induct arbitrary: xs set: Wellfounded.accp)
   apply (erule thin_rl)
   apply (erule accp_induct)
   apply (rule accp.accI)
   apply blast
   done
 
-lemma lists_accD: "listsp (accp r) xs ==> accp (step1 r) xs"
+lemma lists_accD: "listsp (Wellfounded.accp r) xs ==> Wellfounded.accp (step1 r) xs"
   apply (induct set: listsp)
    apply (rule accp.accI)
    apply simp
@@ -113,8 +113,8 @@
   apply force
   done
 
-lemma lists_accI: "accp (step1 r) xs ==> listsp (accp r) xs"
-  apply (induct set: accp)
+lemma lists_accI: "Wellfounded.accp (step1 r) xs ==> listsp (Wellfounded.accp r) xs"
+  apply (induct set: Wellfounded.accp)
   apply clarify
   apply (rule accp.accI)
   apply (drule_tac r=r in ex_step1I, assumption)
--- a/src/HOL/ROOT	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/ROOT	Mon Nov 11 17:44:21 2013 +0100
@@ -733,6 +733,7 @@
   theories [condition = ISABELLE_FULL_TEST]
     Misc_Codatatype
     Misc_Datatype
+    Misc_Primcorec
     Misc_Primrec
 
 session "HOL-Word" (main) in Word = HOL +
--- a/src/HOL/Rat.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Rat.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -468,7 +468,7 @@
     unfolding less_eq_rat_def less_rat_def
     by (auto, drule (1) positive_add, simp add: positive_zero)
   show "a \<le> b \<Longrightarrow> c + a \<le> c + b"
-    unfolding less_eq_rat_def less_rat_def by (auto simp: diff_minus)
+    unfolding less_eq_rat_def less_rat_def by auto
   show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
     by (rule sgn_rat_def)
   show "a \<le> b \<or> b \<le> a"
@@ -665,7 +665,7 @@
   by transfer simp
 
 lemma of_rat_diff: "of_rat (a - b) = of_rat a - of_rat b"
-by (simp only: diff_minus of_rat_add of_rat_minus)
+  using of_rat_add [of a "- b"] by (simp add: of_rat_minus)
 
 lemma of_rat_mult: "of_rat (a * b) = of_rat a * of_rat b"
 apply transfer
--- a/src/HOL/Real.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Real.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -98,7 +98,7 @@
 lemma vanishes_diff:
   assumes X: "vanishes X" and Y: "vanishes Y"
   shows "vanishes (\<lambda>n. X n - Y n)"
-unfolding diff_minus by (intro vanishes_add vanishes_minus X Y)
+  unfolding diff_conv_add_uminus by (intro vanishes_add vanishes_minus X Y)
 
 lemma vanishes_mult_bounded:
   assumes X: "\<exists>a>0. \<forall>n. \<bar>X n\<bar> < a"
@@ -170,7 +170,7 @@
 lemma cauchy_diff [simp]:
   assumes X: "cauchy X" and Y: "cauchy Y"
   shows "cauchy (\<lambda>n. X n - Y n)"
-using assms unfolding diff_minus by simp
+  using assms unfolding diff_conv_add_uminus by (simp del: add_uminus_conv_diff)
 
 lemma cauchy_imp_bounded:
   assumes "cauchy X" shows "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
@@ -456,7 +456,7 @@
 lemma diff_Real:
   assumes X: "cauchy X" and Y: "cauchy Y"
   shows "Real X - Real Y = Real (\<lambda>n. X n - Y n)"
-  unfolding minus_real_def diff_minus
+  unfolding minus_real_def
   by (simp add: minus_Real add_Real X Y)
 
 lemma mult_Real:
@@ -607,7 +607,7 @@
     unfolding less_eq_real_def less_real_def
     by (auto, drule (1) positive_add, simp add: positive_zero)
   show "a \<le> b \<Longrightarrow> c + a \<le> c + b"
-    unfolding less_eq_real_def less_real_def by (auto simp: diff_minus) (* by auto *)
+    unfolding less_eq_real_def less_real_def by auto
     (* FIXME: Procedure int_combine_numerals: c + b - (c + a) \<equiv> b + - a *)
     (* Should produce c + b - (c + a) \<equiv> b - a *)
   show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
@@ -919,24 +919,20 @@
     using 1 2 3 by (rule_tac x="Real B" in exI, simp)
 qed
 
-
 instantiation real :: linear_continuum
 begin
 
 subsection{*Supremum of a set of reals*}
 
-definition
-  Sup_real_def: "Sup X \<equiv> LEAST z::real. \<forall>x\<in>X. x\<le>z"
-
-definition
-  Inf_real_def: "Inf (X::real set) \<equiv> - Sup (uminus ` X)"
+definition "Sup X = (LEAST z::real. \<forall>x\<in>X. x \<le> z)"
+definition "Inf (X::real set) = - Sup (uminus ` X)"
 
 instance
 proof
-  { fix z x :: real and X :: "real set"
-    assume x: "x \<in> X" and z: "\<And>x. x \<in> X \<Longrightarrow> x \<le> z"
+  { fix x :: real and X :: "real set"
+    assume x: "x \<in> X" "bdd_above X"
     then obtain s where s: "\<forall>y\<in>X. y \<le> s" "\<And>z. \<forall>y\<in>X. y \<le> z \<Longrightarrow> s \<le> z"
-      using complete_real[of X] by blast
+      using complete_real[of X] unfolding bdd_above_def by blast
     then show "x \<le> Sup X"
       unfolding Sup_real_def by (rule LeastI2_order) (auto simp: x) }
   note Sup_upper = this
@@ -952,32 +948,15 @@
     finally show "Sup X \<le> z" . }
   note Sup_least = this
 
-  { fix x z :: real and X :: "real set"
-    assume x: "x \<in> X" and z: "\<And>x. x \<in> X \<Longrightarrow> z \<le> x"
-    have "-x \<le> Sup (uminus ` X)"
-      by (rule Sup_upper[of _ _ "- z"]) (auto simp add: image_iff x z)
-    then show "Inf X \<le> x" 
-      by (auto simp add: Inf_real_def) }
-
-  { fix z :: real and X :: "real set"
-    assume x: "X \<noteq> {}" and z: "\<And>x. x \<in> X \<Longrightarrow> z \<le> x"
-    have "Sup (uminus ` X) \<le> -z"
-      using x z by (force intro: Sup_least)
-    then show "z \<le> Inf X" 
-        by (auto simp add: Inf_real_def) }
-
+  { fix x :: real and X :: "real set" assume x: "x \<in> X" "bdd_below X" then show "Inf X \<le> x"
+      using Sup_upper[of "-x" "uminus ` X"] by (auto simp: Inf_real_def) }
+  { fix z :: real and X :: "real set" assume "X \<noteq> {}" "\<And>x. x \<in> X \<Longrightarrow> z \<le> x" then show "z \<le> Inf X"
+      using Sup_least[of "uminus ` X" "- z"] by (force simp: Inf_real_def) }
   show "\<exists>a b::real. a \<noteq> b"
     using zero_neq_one by blast
 qed
 end
 
-text {*
-  \medskip Completeness properties using @{text "isUb"}, @{text "isLub"}:
-*}
-
-lemma reals_complete: "\<exists>X. X \<in> S \<Longrightarrow> \<exists>Y. isUb (UNIV::real set) S Y \<Longrightarrow> \<exists>t. isLub (UNIV :: real set) S t"
-  by (intro exI[of _ "Sup S"] isLub_cSup) (auto simp: setle_def isUb_def intro: cSup_upper)
-
 
 subsection {* Hiding implementation details *}
 
--- a/src/HOL/Real_Vector_Spaces.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Real_Vector_Spaces.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -31,7 +31,7 @@
 qed
 
 lemma diff: "f (x - y) = f x - f y"
-by (simp add: add minus diff_minus)
+  using add [of x "- y"] by (simp add: minus)
 
 lemma setsum: "f (setsum g A) = (\<Sum>x\<in>A. f (g x))"
 apply (cases "finite A")
@@ -553,8 +553,7 @@
 proof -
   have "norm (a + - b) \<le> norm a + norm (- b)"
     by (rule norm_triangle_ineq)
-  thus ?thesis
-    by (simp only: diff_minus norm_minus_cancel)
+  then show ?thesis by simp
 qed
 
 lemma norm_diff_ineq:
@@ -571,7 +570,7 @@
   shows "norm ((a + b) - (c + d)) \<le> norm (a - c) + norm (b - d)"
 proof -
   have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))"
-    by (simp add: diff_minus add_ac)
+    by (simp add: algebra_simps)
   also have "\<dots> \<le> norm (a - c) + norm (b - d)"
     by (rule norm_triangle_ineq)
   finally show ?thesis .
@@ -1426,9 +1425,6 @@
   @{term "{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}"}
 *}
 
-lemma isUb_UNIV_I: "(\<And>y. y \<in> S \<Longrightarrow> y \<le> u) \<Longrightarrow> isUb UNIV S u"
-by (simp add: isUbI setleI)
-
 lemma increasing_LIMSEQ:
   fixes f :: "nat \<Rightarrow> real"
   assumes inc: "\<And>n. f n \<le> f (Suc n)"
@@ -1455,40 +1451,33 @@
   then have mem_S: "\<And>N x. \<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S" by auto
 
   { fix N x assume N: "\<forall>n\<ge>N. X n < x"
-  have "isUb UNIV S x"
-  proof (rule isUb_UNIV_I)
   fix y::real assume "y \<in> S"
   hence "\<exists>M. \<forall>n\<ge>M. y < X n"
     by (simp add: S_def)
   then obtain M where "\<forall>n\<ge>M. y < X n" ..
   hence "y < X (max M N)" by simp
   also have "\<dots> < x" using N by simp
-  finally show "y \<le> x"
-    by (rule order_less_imp_le)
-  qed }
+  finally have "y \<le> x"
+    by (rule order_less_imp_le) }
   note bound_isUb = this 
 
-  have "\<exists>u. isLub UNIV S u"
-  proof (rule reals_complete)
   obtain N where "\<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m) (X n) < 1"
     using X[THEN metric_CauchyD, OF zero_less_one] by auto
   hence N: "\<forall>n\<ge>N. dist (X n) (X N) < 1" by simp
-  show "\<exists>x. x \<in> S"
-  proof
+  have [simp]: "S \<noteq> {}"
+  proof (intro exI ex_in_conv[THEN iffD1])
     from N have "\<forall>n\<ge>N. X N - 1 < X n"
       by (simp add: abs_diff_less_iff dist_real_def)
     thus "X N - 1 \<in> S" by (rule mem_S)
   qed
-  show "\<exists>u. isUb UNIV S u"
+  have [simp]: "bdd_above S"
   proof
     from N have "\<forall>n\<ge>N. X n < X N + 1"
       by (simp add: abs_diff_less_iff dist_real_def)
-    thus "isUb UNIV S (X N + 1)"
+    thus "\<And>s. s \<in> S \<Longrightarrow>  s \<le> X N + 1"
       by (rule bound_isUb)
   qed
-  qed
-  then obtain x where x: "isLub UNIV S x" ..
-  have "X ----> x"
+  have "X ----> Sup S"
   proof (rule metric_LIMSEQ_I)
   fix r::real assume "0 < r"
   hence r: "0 < r/2" by simp
@@ -1500,17 +1489,18 @@
 
   from N have "\<forall>n\<ge>N. X N - r/2 < X n" by fast
   hence "X N - r/2 \<in> S" by (rule mem_S)
-  hence 1: "X N - r/2 \<le> x" using x isLub_isUb isUbD by fast
+  hence 1: "X N - r/2 \<le> Sup S" by (simp add: cSup_upper)
 
   from N have "\<forall>n\<ge>N. X n < X N + r/2" by fast
-  hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb)
-  hence 2: "x \<le> X N + r/2" using x isLub_le_isUb by fast
+  from bound_isUb[OF this]
+  have 2: "Sup S \<le> X N + r/2"
+    by (intro cSup_least) simp_all
 
-  show "\<exists>N. \<forall>n\<ge>N. dist (X n) x < r"
+  show "\<exists>N. \<forall>n\<ge>N. dist (X n) (Sup S) < r"
   proof (intro exI allI impI)
     fix n assume n: "N \<le> n"
     from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+
-    thus "dist (X n) x < r" using 1 2
+    thus "dist (X n) (Sup S) < r" using 1 2
       by (simp add: abs_diff_less_iff dist_real_def)
   qed
   qed
--- a/src/HOL/Record.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Record.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -399,7 +399,7 @@
 lemma refl_conj_eq: "Q = R \<Longrightarrow> P \<and> Q \<longleftrightarrow> P \<and> R"
   by simp
 
-lemma iso_tuple_UNIV_I [no_atp]: "x \<in> UNIV \<equiv> True"
+lemma iso_tuple_UNIV_I: "x \<in> UNIV \<equiv> True"
   by simp
 
 lemma iso_tuple_True_simp: "(True \<Longrightarrow> PROP P) \<equiv> PROP P"
--- a/src/HOL/Relation.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Relation.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -478,7 +478,7 @@
 lemma Id_on_eqI: "a = b ==> a : A ==> (a, b) : Id_on A"
   by (simp add: Id_on_def)
 
-lemma Id_onI [intro!,no_atp]: "a : A ==> (a, a) : Id_on A"
+lemma Id_onI [intro!]: "a : A ==> (a, a) : Id_on A"
   by (rule Id_on_eqI) (rule refl)
 
 lemma Id_onE [elim!]:
@@ -939,8 +939,6 @@
 where
   "r `` s = {y. \<exists>x\<in>s. (x, y) \<in> r}"
 
-declare Image_def [no_atp]
-
 lemma Image_iff: "(b : r``A) = (EX x:A. (x, b) : r)"
   by (simp add: Image_def)
 
@@ -950,7 +948,7 @@
 lemma Image_singleton_iff [iff]: "(b : r``{a}) = ((a, b) : r)"
   by (rule Image_iff [THEN trans]) simp
 
-lemma ImageI [intro,no_atp]: "(a, b) : r ==> a : A ==> b : r``A"
+lemma ImageI [intro]: "(a, b) : r ==> a : A ==> b : r``A"
   by (unfold Image_def) blast
 
 lemma ImageE [elim!]:
--- a/src/HOL/Rings.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Rings.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -86,7 +86,20 @@
 lemma one_neq_zero [simp]: "1 \<noteq> 0"
 by (rule not_sym) (rule zero_neq_one)
 
-end
+definition of_bool :: "bool \<Rightarrow> 'a"
+where
+  "of_bool p = (if p then 1 else 0)" 
+
+lemma of_bool_eq [simp, code]:
+  "of_bool False = 0"
+  "of_bool True = 1"
+  by (simp_all add: of_bool_def)
+
+lemma of_bool_eq_iff:
+  "of_bool p = of_bool q \<longleftrightarrow> p = q"
+  by (simp add: of_bool_def)
+
+end  
 
 class semiring_1 = zero_neq_one + semiring_0 + monoid_mult
 
@@ -127,7 +140,7 @@
   then show ?thesis ..
 qed
 
-lemma dvd_0_left_iff [no_atp, simp]: "0 dvd a \<longleftrightarrow> a = 0"
+lemma dvd_0_left_iff [simp]: "0 dvd a \<longleftrightarrow> a = 0"
 by (auto intro: dvd_refl elim!: dvdE)
 
 lemma dvd_0_right [iff]: "a dvd 0"
@@ -233,8 +246,8 @@
 by (rule minus_unique) (simp add: distrib_left [symmetric]) 
 
 text{*Extract signs from products*}
-lemmas mult_minus_left [simp, no_atp] = minus_mult_left [symmetric]
-lemmas mult_minus_right [simp,no_atp] = minus_mult_right [symmetric]
+lemmas mult_minus_left [simp] = minus_mult_left [symmetric]
+lemmas mult_minus_right [simp] = minus_mult_right [symmetric]
 
 lemma minus_mult_minus [simp]: "- a * - b = a * b"
 by simp
@@ -242,13 +255,15 @@
 lemma minus_mult_commute: "- a * b = a * - b"
 by simp
 
-lemma right_diff_distrib[algebra_simps, field_simps]: "a * (b - c) = a * b - a * c"
-by (simp add: distrib_left diff_minus)
+lemma right_diff_distrib [algebra_simps, field_simps]:
+  "a * (b - c) = a * b - a * c"
+  using distrib_left [of a b "-c "] by simp
 
-lemma left_diff_distrib[algebra_simps, field_simps]: "(a - b) * c = a * c - b * c"
-by (simp add: distrib_right diff_minus)
+lemma left_diff_distrib [algebra_simps, field_simps]:
+  "(a - b) * c = a * c - b * c"
+  using distrib_right [of a "- b" c] by simp
 
-lemmas ring_distribs[no_atp] =
+lemmas ring_distribs =
   distrib_left distrib_right left_diff_distrib right_diff_distrib
 
 lemma eq_add_iff1:
@@ -261,7 +276,7 @@
 
 end
 
-lemmas ring_distribs[no_atp] =
+lemmas ring_distribs =
   distrib_left distrib_right left_diff_distrib right_diff_distrib
 
 class comm_ring = comm_semiring + ab_group_add
@@ -318,8 +333,9 @@
   then show "- x dvd y" ..
 qed
 
-lemma dvd_diff[simp]: "x dvd y \<Longrightarrow> x dvd z \<Longrightarrow> x dvd (y - z)"
-by (simp only: diff_minus dvd_add dvd_minus_iff)
+lemma dvd_diff [simp]:
+  "x dvd y \<Longrightarrow> x dvd z \<Longrightarrow> x dvd (y - z)"
+  using dvd_add [of x y "- z"] by simp
 
 end
 
@@ -336,7 +352,7 @@
 qed
 
 text{*Cancellation of equalities with a common factor*}
-lemma mult_cancel_right [simp, no_atp]:
+lemma mult_cancel_right [simp]:
   "a * c = b * c \<longleftrightarrow> c = 0 \<or> a = b"
 proof -
   have "(a * c = b * c) = ((a - b) * c = 0)"
@@ -344,7 +360,7 @@
   thus ?thesis by (simp add: disj_commute)
 qed
 
-lemma mult_cancel_left [simp, no_atp]:
+lemma mult_cancel_left [simp]:
   "c * a = c * b \<longleftrightarrow> c = 0 \<or> a = b"
 proof -
   have "(c * a = c * b) = (c * (a - b) = 0)"
@@ -742,9 +758,7 @@
 proof
   fix a b
   show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
-    by (auto simp add: abs_if not_less)
-    (auto simp del: minus_add_distrib simp add: minus_add_distrib [symmetric],
-     auto intro!: less_imp_le add_neg_neg)
+    by (auto simp add: abs_if not_le not_less algebra_simps simp del: add.commute dest: add_neg_neg add_nonneg_nonneg)
 qed (auto simp add: abs_if)
 
 lemma zero_le_square [simp]: "0 \<le> a * a"
@@ -1048,7 +1062,7 @@
 
 text {* Simprules for comparisons where common factors can be cancelled. *}
 
-lemmas mult_compare_simps[no_atp] =
+lemmas mult_compare_simps =
     mult_le_cancel_right mult_le_cancel_left
     mult_le_cancel_right1 mult_le_cancel_right2
     mult_le_cancel_left1 mult_le_cancel_left2
@@ -1131,10 +1145,6 @@
   thus ?thesis by (simp add: ac cpos mult_strict_mono) 
 qed
 
-lemma less_minus_self_iff:
-  "a < - a \<longleftrightarrow> a < 0"
-  by (simp only: less_le less_eq_neg_nonpos equal_neg_zero)
-
 lemma abs_less_iff:
   "\<bar>a\<bar> < b \<longleftrightarrow> a < b \<and> - a < b" 
   by (simp add: less_le abs_le_iff) (auto simp add: abs_if)
--- a/src/HOL/Semiring_Normalization.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Semiring_Normalization.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -137,7 +137,7 @@
 lemma normalizing_ring_rules:
   "- x = (- 1) * x"
   "x - y = x + (- y)"
-  by (simp_all add: diff_minus)
+  by simp_all
 
 lemmas normalizing_comm_ring_1_axioms =
   comm_ring_1_axioms [normalizer
--- a/src/HOL/Series.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Series.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -34,7 +34,7 @@
 
 lemma sumr_diff_mult_const:
  "setsum f {0..<n} - (real n*r) = setsum (%i. f i - r) {0..<n::nat}"
-by (simp add: diff_minus setsum_addf real_of_nat_def)
+  by (simp add: setsum_subtractf real_of_nat_def)
 
 lemma real_setsum_nat_ivl_bounded:
      "(!!p. p < n \<Longrightarrow> f(p) \<le> K)
--- a/src/HOL/Set.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Set.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -87,7 +87,7 @@
   then show ?thesis by simp
 qed
 
-lemma set_eq_iff [no_atp]:
+lemma set_eq_iff:
   "A = B \<longleftrightarrow> (\<forall>x. x \<in> A \<longleftrightarrow> x \<in> B)"
   by (auto intro:set_eqI)
 
@@ -495,7 +495,7 @@
   by (simp add: less_eq_set_def le_fun_def)
   -- {* Rule in Modus Ponens style. *}
 
-lemma rev_subsetD [no_atp,intro?]: "c \<in> A ==> A \<subseteq> B ==> c \<in> B"
+lemma rev_subsetD [intro?]: "c \<in> A ==> A \<subseteq> B ==> c \<in> B"
   -- {* The same, with reversed premises for use with @{text erule} --
       cf @{text rev_mp}. *}
   by (rule subsetD)
@@ -504,13 +504,13 @@
   \medskip Converts @{prop "A \<subseteq> B"} to @{prop "x \<in> A ==> x \<in> B"}.
 *}
 
-lemma subsetCE [no_atp,elim]: "A \<subseteq> B ==> (c \<notin> A ==> P) ==> (c \<in> B ==> P) ==> P"
+lemma subsetCE [elim]: "A \<subseteq> B ==> (c \<notin> A ==> P) ==> (c \<in> B ==> P) ==> P"
   -- {* Classical elimination rule. *}
   by (auto simp add: less_eq_set_def le_fun_def)
 
-lemma subset_eq [no_atp]: "A \<le> B = (\<forall>x\<in>A. x \<in> B)" by blast
-
-lemma contra_subsetD [no_atp]: "A \<subseteq> B ==> c \<notin> B ==> c \<notin> A"
+lemma subset_eq: "A \<le> B = (\<forall>x\<in>A. x \<in> B)" by blast
+
+lemma contra_subsetD: "A \<subseteq> B ==> c \<notin> B ==> c \<notin> A"
   by blast
 
 lemma subset_refl: "A \<subseteq> A"
@@ -845,11 +845,11 @@
 
 subsubsection {* Singletons, using insert *}
 
-lemma singletonI [intro!,no_atp]: "a : {a}"
+lemma singletonI [intro!]: "a : {a}"
     -- {* Redundant? But unlike @{text insertCI}, it proves the subgoal immediately! *}
   by (rule insertI1)
 
-lemma singletonD [dest!,no_atp]: "b : {a} ==> b = a"
+lemma singletonD [dest!]: "b : {a} ==> b = a"
   by blast
 
 lemmas singletonE = singletonD [elim_format]
@@ -860,11 +860,11 @@
 lemma singleton_inject [dest!]: "{a} = {b} ==> a = b"
   by blast
 
-lemma singleton_insert_inj_eq [iff,no_atp]:
+lemma singleton_insert_inj_eq [iff]:
      "({b} = insert a A) = (a = b & A \<subseteq> {b})"
   by blast
 
-lemma singleton_insert_inj_eq' [iff,no_atp]:
+lemma singleton_insert_inj_eq' [iff]:
      "(insert a A = {b}) = (a = b & A \<subseteq> {b})"
   by blast
 
@@ -898,7 +898,7 @@
 *}
 
 definition image :: "('a => 'b) => 'a set => 'b set" (infixr "`" 90) where
-  image_def [no_atp]: "f ` A = {y. EX x:A. y = f(x)}"
+  image_def: "f ` A = {y. EX x:A. y = f(x)}"
 
 abbreviation
   range :: "('a => 'b) => 'b set" where -- "of function"
@@ -930,7 +930,7 @@
 lemma image_iff: "(z : f`A) = (EX x:A. z = f x)"
   by blast
 
-lemma image_subset_iff [no_atp]: "(f`A \<subseteq> B) = (\<forall>x\<in>A. f x \<in> B)"
+lemma image_subset_iff: "(f`A \<subseteq> B) = (\<forall>x\<in>A. f x \<in> B)"
   -- {* This rewrite rule would confuse users if made default. *}
   by blast
 
@@ -1009,10 +1009,10 @@
 
 subsubsection {* The ``proper subset'' relation *}
 
-lemma psubsetI [intro!,no_atp]: "A \<subseteq> B ==> A \<noteq> B ==> A \<subset> B"
+lemma psubsetI [intro!]: "A \<subseteq> B ==> A \<noteq> B ==> A \<subset> B"
   by (unfold less_le) blast
 
-lemma psubsetE [elim!,no_atp]:
+lemma psubsetE [elim!]:
     "[|A \<subset> B;  [|A \<subseteq> B; ~ (B\<subseteq>A)|] ==> R|] ==> R"
   by (unfold less_le) blast
 
@@ -1184,12 +1184,12 @@
 lemma insert_inter_insert[simp]: "insert a A \<inter> insert a B = insert a (A \<inter> B)"
   by blast
 
-lemma insert_disjoint [simp,no_atp]:
+lemma insert_disjoint [simp]:
  "(insert a A \<inter> B = {}) = (a \<notin> B \<and> A \<inter> B = {})"
  "({} = insert a A \<inter> B) = (a \<notin> B \<and> {} = A \<inter> B)"
   by auto
 
-lemma disjoint_insert [simp,no_atp]:
+lemma disjoint_insert [simp]:
  "(B \<inter> insert a A = {}) = (a \<notin> B \<and> B \<inter> A = {})"
  "({} = A \<inter> insert b B) = (b \<notin> A \<and> {} = A \<inter> B)"
   by auto
@@ -1221,7 +1221,7 @@
 by blast
 
 
-lemma image_Collect [no_atp]: "f ` {x. P x} = {f x | x. P x}"
+lemma image_Collect: "f ` {x. P x} = {f x | x. P x}"
   -- {* NOT suitable as a default simprule: the RHS isn't simpler than the LHS,
       with its implicit quantifier and conjunction.  Also image enjoys better
       equational properties than does the RHS. *}
@@ -1244,7 +1244,7 @@
 
 text {* \medskip @{text range}. *}
 
-lemma full_SetCompr_eq [no_atp]: "{u. \<exists>x. u = f x} = range f"
+lemma full_SetCompr_eq: "{u. \<exists>x. u = f x} = range f"
   by auto
 
 lemma range_composition: "range (\<lambda>x. f (g x)) = f`range g"
@@ -1301,10 +1301,10 @@
 lemma Int_Un_distrib2: "(B \<union> C) \<inter> A = (B \<inter> A) \<union> (C \<inter> A)"
   by (fact inf_sup_distrib2)
 
-lemma Int_UNIV [simp,no_atp]: "(A \<inter> B = UNIV) = (A = UNIV & B = UNIV)"
+lemma Int_UNIV [simp]: "(A \<inter> B = UNIV) = (A = UNIV & B = UNIV)"
   by (fact inf_eq_top_iff) (* already simp *)
 
-lemma Int_subset_iff [no_atp, simp]: "(C \<subseteq> A \<inter> B) = (C \<subseteq> A & C \<subseteq> B)"
+lemma Int_subset_iff [simp]: "(C \<subseteq> A \<inter> B) = (C \<subseteq> A & C \<subseteq> B)"
   by (fact le_inf_iff)
 
 lemma Int_Collect: "(x \<in> A \<inter> {x. P x}) = (x \<in> A & P x)"
@@ -1395,7 +1395,7 @@
 lemma Un_empty [iff]: "(A \<union> B = {}) = (A = {} & B = {})"
   by (fact sup_eq_bot_iff) (* FIXME: already simp *)
 
-lemma Un_subset_iff [no_atp, simp]: "(A \<union> B \<subseteq> C) = (A \<subseteq> C & B \<subseteq> C)"
+lemma Un_subset_iff [simp]: "(A \<union> B \<subseteq> C) = (A \<subseteq> C & B \<subseteq> C)"
   by (fact le_sup_iff)
 
 lemma Un_Diff_Int: "(A - B) \<union> (A \<inter> B) = A"
@@ -1467,7 +1467,7 @@
 lemma Diff_eq: "A - B = A \<inter> (-B)"
   by blast
 
-lemma Diff_eq_empty_iff [simp,no_atp]: "(A - B = {}) = (A \<subseteq> B)"
+lemma Diff_eq_empty_iff [simp]: "(A - B = {}) = (A \<subseteq> B)"
   by blast
 
 lemma Diff_cancel [simp]: "A - A = {}"
@@ -1488,7 +1488,7 @@
 lemma Diff_UNIV [simp]: "A - UNIV = {}"
   by blast
 
-lemma Diff_insert0 [simp,no_atp]: "x \<notin> A ==> A - insert x B = A - B"
+lemma Diff_insert0 [simp]: "x \<notin> A ==> A - insert x B = A - B"
   by blast
 
 lemma Diff_insert: "A - insert a B = A - B - {a}"
@@ -1568,7 +1568,7 @@
 lemma ex_bool_eq: "(\<exists>b. P b) \<longleftrightarrow> P True \<or> P False"
   by (auto intro: bool_contrapos)
 
-lemma UNIV_bool [no_atp]: "UNIV = {False, True}"
+lemma UNIV_bool: "UNIV = {False, True}"
   by (auto intro: bool_induct)
 
 text {* \medskip @{text Pow} *}
@@ -1597,7 +1597,7 @@
 lemma set_eq_subset: "(A = B) = (A \<subseteq> B & B \<subseteq> A)"
   by blast
 
-lemma subset_iff [no_atp]: "(A \<subseteq> B) = (\<forall>t. t \<in> A --> t \<in> B)"
+lemma subset_iff: "(A \<subseteq> B) = (\<forall>t. t \<in> A --> t \<in> B)"
   by blast
 
 lemma subset_iff_psubset_eq: "(A \<subseteq> B) = ((A \<subset> B) | (A = B))"
@@ -1754,7 +1754,7 @@
   -- {* monotonicity *}
   by blast
 
-lemma vimage_image_eq [no_atp]: "f -` (f ` A) = {y. EX x:A. f x = f y}"
+lemma vimage_image_eq: "f -` (f ` A) = {y. EX x:A. f x = f y}"
 by (blast intro: sym)
 
 lemma image_vimage_subset: "f ` (f -` A) <= A"
--- a/src/HOL/Set_Interval.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Set_Interval.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -180,19 +180,19 @@
 context ord
 begin
 
-lemma greaterThanLessThan_iff [simp,no_atp]:
+lemma greaterThanLessThan_iff [simp]:
   "(i : {l<..<u}) = (l < i & i < u)"
 by (simp add: greaterThanLessThan_def)
 
-lemma atLeastLessThan_iff [simp,no_atp]:
+lemma atLeastLessThan_iff [simp]:
   "(i : {l..<u}) = (l <= i & i < u)"
 by (simp add: atLeastLessThan_def)
 
-lemma greaterThanAtMost_iff [simp,no_atp]:
+lemma greaterThanAtMost_iff [simp]:
   "(i : {l<..u}) = (l < i & i <= u)"
 by (simp add: greaterThanAtMost_def)
 
-lemma atLeastAtMost_iff [simp,no_atp]:
+lemma atLeastAtMost_iff [simp]:
   "(i : {l..u}) = (l <= i & i <= u)"
 by (simp add: atLeastAtMost_def)
 
@@ -1196,7 +1196,7 @@
 
 subsubsection {* Some Subset Conditions *}
 
-lemma ivl_subset [simp,no_atp]:
+lemma ivl_subset [simp]:
  "({i..<j} \<subseteq> {m..<n}) = (j \<le> i | m \<le> i & j \<le> (n::'a::linorder))"
 apply(auto simp:linorder_not_le)
 apply(rule ccontr)
--- a/src/HOL/TPTP/MaSh_Export.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/TPTP/MaSh_Export.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -63,7 +63,7 @@
 
 ML {*
 if do_it then
-  generate_isar_dependencies @{context} thys linearize
+  generate_isar_dependencies @{context} range thys linearize
       (prefix ^ "mash_dependencies")
 else
   ()
--- a/src/HOL/TPTP/atp_theory_export.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/TPTP/atp_theory_export.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -34,8 +34,8 @@
 val prefix = Library.prefix
 val fact_name_of = prefix fact_prefix o ascii_of
 
-fun atp_of_format (THF (Polymorphic, _, _)) = dummy_thfN
-  | atp_of_format (THF (Monomorphic, _, _)) = satallaxN
+fun atp_of_format (THF (Polymorphic, _)) = dummy_thfN
+  | atp_of_format (THF (Monomorphic, _)) = satallaxN
   | atp_of_format (DFG Polymorphic) = spass_polyN
   | atp_of_format (DFG Monomorphic) = spassN
   | atp_of_format (TFF Polymorphic) = alt_ergoN
--- a/src/HOL/TPTP/mash_eval.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/TPTP/mash_eval.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -147,7 +147,7 @@
                   |> map fact_of_raw_fact
                 val ctxt = ctxt |> set_file_name method prob_dir_name
                 val res as {outcome, ...} =
-                  run_prover_for_mash ctxt params prover facts goal
+                  run_prover_for_mash ctxt params prover name facts goal
                 val ok = if is_none outcome then 1 else 0
               in (str_of_result method facts res, ok) end
           val ress =
--- a/src/HOL/TPTP/mash_export.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/TPTP/mash_export.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -14,7 +14,7 @@
   val generate_features :
     Proof.context -> string -> theory list -> string -> unit
   val generate_isar_dependencies :
-    Proof.context -> theory list -> bool -> string -> unit
+    Proof.context -> int * int option -> theory list -> bool -> string -> unit
   val generate_prover_dependencies :
     Proof.context -> params -> int * int option -> theory list -> bool -> string
     -> unit
@@ -79,9 +79,7 @@
       let
         val name = nickname_of_thm th
         val feats =
-          features_of ctxt prover (theory_of_thm th) 0 Symtab.empty stature
-                      [prop_of th]
-          |> map fst
+          features_of ctxt (theory_of_thm th) 0 Symtab.empty stature [prop_of th] |> map fst
         val s =
           encode_str name ^ ": " ^ encode_strs (sort string_ord feats) ^ "\n"
       in File.append path s end
@@ -136,7 +134,7 @@
   in File.write_list path lines end
 
 fun generate_isar_dependencies ctxt =
-  generate_isar_or_prover_dependencies ctxt NONE (1, NONE)
+  generate_isar_or_prover_dependencies ctxt NONE
 
 fun generate_prover_dependencies ctxt params =
   generate_isar_or_prover_dependencies ctxt (SOME params)
@@ -164,8 +162,7 @@
           val isar_deps = isar_dependencies_of name_tabs th
           val do_query = not (is_bad_query ctxt ho_atp step j th isar_deps)
           val goal_feats =
-            features_of ctxt prover (theory_of_thm th) (num_old_facts + j)
-                        const_tab stature [prop_of th]
+            features_of ctxt (theory_of_thm th) (num_old_facts + j) const_tab stature [prop_of th]
             |> sort_wrt fst
           val access_facts =
             (if linearize then take (j - 1) new_facts
@@ -176,8 +173,7 @@
           val parents = if linearize then prevs else parents
           fun extra_features_of (((_, stature), th), weight) =
             [prop_of th]
-            |> features_of ctxt prover (theory_of_thm th) (num_old_facts + j)
-                           const_tab stature
+            |> features_of ctxt (theory_of_thm th) (num_old_facts + j) const_tab stature
             |> map (apsnd (fn r => weight * extra_feature_factor * r))
           val query =
             if do_query then
@@ -249,8 +245,8 @@
               val suggs =
                 old_facts
                 |> linearize ? filter_accessible_from th
-                |> Sledgehammer_MePo.mepo_suggested_facts ctxt params prover
-                       max_suggs NONE hyp_ts concl_t
+                |> Sledgehammer_Fact.drop_duplicate_facts
+                |> Sledgehammer_MePo.mepo_suggested_facts ctxt params max_suggs NONE hyp_ts concl_t
                 |> map (nickname_of_thm o snd)
             in encode_str name ^ ": " ^ encode_strs suggs ^ "\n" end
         end
--- a/src/HOL/TPTP/sledgehammer_tactics.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/TPTP/sledgehammer_tactics.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -27,24 +27,22 @@
 fun run_prover override_params fact_override i n ctxt goal =
   let
     val mode = Normal
-    val params as {provers, max_facts, slice, ...} =
-      default_params ctxt override_params
+    val params as {provers, max_facts, ...} = default_params ctxt override_params
     val name = hd provers
     val prover = get_prover ctxt mode name
-    val default_max_facts = default_max_facts_of_prover ctxt slice name
+    val default_max_facts = default_max_facts_of_prover ctxt name
     val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal goal i ctxt
     val ho_atp = exists (is_ho_atp ctxt) provers
     val reserved = reserved_isar_keyword_table ()
     val css_table = clasimpset_rule_table_of ctxt
     val facts =
-      nearly_all_facts ctxt ho_atp fact_override reserved css_table [] hyp_ts
-                       concl_t
+      nearly_all_facts ctxt ho_atp fact_override reserved css_table [] hyp_ts concl_t
       |> relevant_facts ctxt params name
              (the_default default_max_facts max_facts) fact_override hyp_ts
              concl_t
       |> hd |> snd
     val problem =
-      {state = Proof.init ctxt, goal = goal, subgoal = i, subgoal_count = n,
+      {comment = "", state = Proof.init ctxt, goal = goal, subgoal = i, subgoal_count = n,
        factss = [("", facts)]}
   in
     (case prover params (K (K (K ""))) problem of
--- a/src/HOL/Tools/ATP/atp_problem.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/ATP/atp_problem.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -33,14 +33,13 @@
 
   datatype polymorphism = Monomorphic | Polymorphic
   datatype thf_choice = THF_Without_Choice | THF_With_Choice
-  datatype thf_defs = THF_Without_Defs | THF_With_Defs
 
   datatype atp_format =
     CNF |
     CNF_UEQ |
     FOF |
     TFF of polymorphism |
-    THF of polymorphism * thf_choice * thf_defs |
+    THF of polymorphism * thf_choice |
     DFG of polymorphism
 
   datatype atp_formula_role =
@@ -179,14 +178,13 @@
 
 datatype polymorphism = Monomorphic | Polymorphic
 datatype thf_choice = THF_Without_Choice | THF_With_Choice
-datatype thf_defs = THF_Without_Defs | THF_With_Defs
 
 datatype atp_format =
   CNF |
   CNF_UEQ |
   FOF |
   TFF of polymorphism |
-  THF of polymorphism * thf_choice * thf_defs |
+  THF of polymorphism * thf_choice |
   DFG of polymorphism
 
 datatype atp_formula_role =
--- a/src/HOL/Tools/ATP/atp_problem_generate.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/ATP/atp_problem_generate.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -91,6 +91,8 @@
   val atp_logical_consts : string list
   val atp_irrelevant_consts : string list
   val atp_widely_irrelevant_consts : string list
+  val is_irrelevant_const : string -> bool
+  val is_widely_irrelevant_const : string -> bool
   val atp_schematic_consts_of : term -> typ list Symtab.table
   val is_type_enc_higher_order : type_enc -> bool
   val is_type_enc_polymorphic : type_enc -> bool
@@ -405,19 +407,25 @@
 (* These are ignored anyway by the relevance filter (unless they appear in
    higher-order places) but not by the monomorphizer. *)
 val atp_logical_consts =
-  [@{const_name all}, @{const_name "==>"}, @{const_name "=="},
-   @{const_name Trueprop}, @{const_name All}, @{const_name Ex},
+  [@{const_name prop}, @{const_name Pure.conjunction}, @{const_name all}, @{const_name "==>"},
+   @{const_name "=="}, @{const_name Trueprop}, @{const_name All}, @{const_name Ex},
    @{const_name Ex1}, @{const_name Ball}, @{const_name Bex}]
 
 (* These are either simplified away by "Meson.presimplify" (most of the time) or
    handled specially via "fFalse", "fTrue", ..., "fequal". *)
 val atp_irrelevant_consts =
-  [@{const_name False}, @{const_name True}, @{const_name Not},
-   @{const_name conj}, @{const_name disj}, @{const_name implies},
-   @{const_name HOL.eq}, @{const_name If}, @{const_name Let}]
+  [@{const_name False}, @{const_name True}, @{const_name Not}, @{const_name conj},
+   @{const_name disj}, @{const_name implies}, @{const_name HOL.eq}, @{const_name If},
+   @{const_name Let}]
 
 val atp_widely_irrelevant_consts = atp_logical_consts @ atp_irrelevant_consts
 
+val atp_irrelevant_const_tab = Symtab.make (map (rpair ()) atp_irrelevant_consts)
+val atp_widely_irrelevant_const_tab = Symtab.make (map (rpair ()) atp_widely_irrelevant_consts)
+
+val is_irrelevant_const = Symtab.defined atp_irrelevant_const_tab
+val is_widely_irrelevant_const = Symtab.defined atp_widely_irrelevant_const_tab
+
 fun add_schematic_const (x as (_, T)) =
   Monomorph.typ_has_tvars T ? Symtab.insert_list (op =) x
 val add_schematic_consts_of =
@@ -689,11 +697,9 @@
     Raw_Polymorphic With_Phantom_Type_Vars
   | no_type_classes poly = poly
 
-fun adjust_type_enc (THF (Polymorphic, choice, _))
-                    (Native (order, poly, level)) =
+fun adjust_type_enc (THF (Polymorphic, choice)) (Native (order, poly, level)) =
     Native (adjust_order choice order, no_type_classes poly, level)
-  | adjust_type_enc (THF (Monomorphic, choice, _))
-                         (Native (order, _, level)) =
+  | adjust_type_enc (THF (Monomorphic, choice)) (Native (order, _, level)) =
     Native (adjust_order choice order, Mangled_Monomorphic, level)
   | adjust_type_enc (TFF Monomorphic) (Native (_, _, level)) =
     Native (First_Order, Mangled_Monomorphic, level)
@@ -1309,7 +1315,7 @@
      atomic_types = atomic_Ts}
   end
 
-fun is_format_with_defs (THF (_, _, THF_With_Defs)) = true
+fun is_format_with_defs (THF _) = true
   | is_format_with_defs _ = false
 
 fun make_fact ctxt format type_enc iff_for_eq
@@ -1557,8 +1563,7 @@
                   end
                 | NONE =>
                   let
-                    val pred_sym = top_level andalso not bool_vars
-                    val ary =
+                    val max_ary =
                       case unprefix_and_unascii const_prefix s of
                         SOME s =>
                         (if String.isSubstring uncurried_alias_sep s then
@@ -1568,15 +1573,16 @@
                            SOME ary0 => Int.min (ary0, ary)
                          | NONE => ary)
                       | NONE => ary
+                    val pred_sym = top_level andalso max_ary = ary andalso not bool_vars
                     val min_ary =
                       case app_op_level of
-                        Min_App_Op => ary
+                        Min_App_Op => max_ary
                       | Full_App_Op_And_Predicator => 0
-                      | _ => fold (consider_var_ary T) fun_var_Ts ary
+                      | _ => fold (consider_var_ary T) fun_var_Ts max_ary
                   in
                     Symtab.update_new (s,
                         {pred_sym = pred_sym, min_ary = min_ary,
-                         max_ary = ary, types = [T], in_conj = conj_fact})
+                         max_ary = max_ary, types = [T], in_conj = conj_fact})
                         sym_tab
                   end)
              end
@@ -2742,12 +2748,10 @@
   else
     NONE
 
-fun ctrss_of_descr descr =
-  map_filter (ctrs_of_datatype descr) descr
+fun ctrss_of_descr descr = map_filter (ctrs_of_datatype descr) descr
 
 fun all_ctrss_of_datatypes thy =
-  Symtab.fold (snd #> #descr #> ctrss_of_descr #> append) (Datatype.get_all thy)
-              []
+  Symtab.fold (snd #> #descr #> ctrss_of_descr #> append) (Datatype.get_all thy) []
 
 val app_op_and_predicator_threshold = 45
 
@@ -2757,35 +2761,26 @@
   let
     val thy = Proof_Context.theory_of ctxt
     val type_enc = type_enc |> adjust_type_enc format
+    val exporter = (mode = Exporter)
+    val completish = (mode = Sledgehammer_Completish)
     (* Forcing explicit applications is expensive for polymorphic encodings,
        because it takes only one existential variable ranging over "'a => 'b" to
        ruin everything. Hence we do it only if there are few facts (which is
        normally the case for "metis" and the minimizer). *)
     val app_op_level =
-      if mode = Sledgehammer_Completish then
+      if completish then
         Full_App_Op_And_Predicator
-      else if length facts + length hyp_ts
-              > app_op_and_predicator_threshold then
-        if is_type_enc_polymorphic type_enc then Min_App_Op
-        else Sufficient_App_Op
+      else if length facts + length hyp_ts >= app_op_and_predicator_threshold then
+        if is_type_enc_polymorphic type_enc then Min_App_Op else Sufficient_App_Op
       else
         Sufficient_App_Op_And_Predicator
-    val exporter = (mode = Exporter)
-    val completish = (mode = Sledgehammer_Completish)
     val lam_trans =
-      if lam_trans = keep_lamsN andalso
-         not (is_type_enc_higher_order type_enc) then
-        liftingN
-      else
-        lam_trans
-    val (fact_names, classes, conjs, facts, subclass_pairs, tcon_clauses,
-         lifted) =
-      translate_formulas ctxt prem_role format type_enc lam_trans preproc hyp_ts
-                         concl_t facts
-    val (_, sym_tab0) =
-      sym_table_of_facts ctxt type_enc app_op_level conjs facts
-    val mono =
-      conjs @ facts |> mononotonicity_info_of_facts ctxt type_enc completish
+      if lam_trans = keep_lamsN andalso not (is_type_enc_higher_order type_enc) then liftingN
+      else lam_trans
+    val (fact_names, classes, conjs, facts, subclass_pairs, tcon_clauses, lifted) =
+      translate_formulas ctxt prem_role format type_enc lam_trans preproc hyp_ts concl_t facts
+    val (_, sym_tab0) = sym_table_of_facts ctxt type_enc app_op_level conjs facts
+    val mono = conjs @ facts |> mononotonicity_info_of_facts ctxt type_enc completish
     val ctrss = all_ctrss_of_datatypes thy
     fun firstorderize in_helper =
       firstorderize_fact thy ctrss type_enc
--- a/src/HOL/Tools/ATP/atp_systems.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/ATP/atp_systems.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -211,7 +211,7 @@
 
 (* agsyHOL *)
 
-val agsyhol_thf0 = THF (Monomorphic, THF_Without_Choice, THF_With_Defs)
+val agsyhol_thf0 = THF (Monomorphic, THF_Without_Choice)
 
 val agsyhol_config : atp_config =
   {exec = K (["AGSYHOL_HOME"], ["agsyHOL"]),
@@ -461,7 +461,7 @@
 
 (* LEO-II supports definitions, but it performs significantly better on our
    benchmarks when they are not used. *)
-val leo2_thf0 = THF (Monomorphic, THF_Without_Choice, THF_Without_Defs)
+val leo2_thf0 = THF (Monomorphic, THF_Without_Choice)
 
 val leo2_config : atp_config =
   {exec = K (["LEO2_HOME"], ["leo.opt", "leo"]),
@@ -488,7 +488,7 @@
 (* Satallax *)
 
 (* Choice is disabled until there is proper reconstruction for it. *)
-val satallax_thf0 = THF (Monomorphic, THF_Without_Choice, THF_With_Defs)
+val satallax_thf0 = THF (Monomorphic, THF_Without_Choice)
 
 val satallax_config : atp_config =
   {exec = K (["SATALLAX_HOME"], ["satallax.opt", "satallax"]),
@@ -653,7 +653,7 @@
    best_max_mono_iters = default_max_mono_iters,
    best_max_new_mono_instances = default_max_new_mono_instances}
 
-val dummy_thf_format = THF (Polymorphic, THF_With_Choice, THF_With_Defs)
+val dummy_thf_format = THF (Polymorphic, THF_With_Choice)
 val dummy_thf_config =
   dummy_config Hypothesis dummy_thf_format "poly_native_higher" false
 val dummy_thf = (dummy_thfN, fn () => dummy_thf_config)
--- a/src/HOL/Tools/Function/function_common.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Function/function_common.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -119,7 +119,7 @@
 
 fun PROFILE msg = if !profile then timeap_msg msg else I
 
-val acc_const_name = @{const_name accp}
+val acc_const_name = @{const_name Wellfounded.accp}
 fun mk_acc domT R =
   Const (acc_const_name, (domT --> domT --> HOLogic.boolT) --> domT --> HOLogic.boolT) $ R 
 
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -1648,8 +1648,7 @@
     fun do_numeral depth Ts mult T t0 t1 =
       (if is_number_type ctxt T then
          let
-           val j = mult * (HOLogic.dest_num t1)
-                   |> T = nat_T ? Integer.max 0
+           val j = mult * HOLogic.dest_num t1
          in
            if j = 1 then
              raise SAME ()
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -280,13 +280,15 @@
             |> map (fn (resultt, (_, prems)) =>
               Logic.list_implies (prems, HOLogic.mk_Trueprop (list_comb (pred, args @ [resultt]))))
           end
-      val intr_ts = maps mk_intros ((funs ~~ preds) ~~ (argss' ~~ rhss))
+      val intr_tss = map mk_intros ((funs ~~ preds) ~~ (argss' ~~ rhss))
       val (intrs, thy') = thy
         |> Sign.add_consts_i
           (map (fn Const (name, T) => (Binding.name (Long_Name.base_name name), T, NoSyn))
            dst_preds)
         |> fold_map Specification.axiom
-            (map (fn t => ((Binding.name ("unnamed_axiom_" ^ serial_string ()), []), t)) intr_ts)
+            (map_index (fn (j, (predname, t)) =>
+                ((Binding.name (Long_Name.base_name predname ^ "_intro_" ^ string_of_int (j + 1)), []), t))
+              (maps (uncurry (map o pair)) (prednames ~~ intr_tss)))
       val specs = map (fn predname => (predname,
           map Drule.export_without_context (filter (Predicate_Compile_Aux.is_intro predname) intrs)))
         dst_prednames
--- a/src/HOL/Tools/Qelim/qelim.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Qelim/qelim.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -37,7 +37,7 @@
      val th = Thm.abstract_rule s x ((conv env' then_conv ncv env') p')
                    |> Drule.arg_cong_rule e
      val th' = simpex_conv (Thm.rhs_of th)
-     val (l,r) = Thm.dest_equals (cprop_of th')
+     val (_, r) = Thm.dest_equals (cprop_of th')
     in if Thm.is_reflexive th' then Thm.transitive th (qcv env (Thm.rhs_of th))
        else Thm.transitive (Thm.transitive th th') (conv env r) end
   | Const(@{const_name Ex},_)$ _ => (Thm.eta_long_conversion then_conv conv env) p
--- a/src/HOL/Tools/SMT/smt_normalize.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/SMT/smt_normalize.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -532,9 +532,7 @@
   *)
 
   fun is_strange_number ctxt (t as Const (@{const_name neg_numeral}, _) $ _) =
-        (case try HOLogic.dest_number t of
-          SOME (_, i) => SMT_Builtin.is_builtin_num ctxt t andalso i < 2
-        | NONE => false)
+        SMT_Builtin.is_builtin_num ctxt t
     | is_strange_number _ _ = false
 
   val pos_num_ss =
--- a/src/HOL/Tools/Sledgehammer/MaSh/etc/settings	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/MaSh/etc/settings	Mon Nov 11 17:44:21 2013 +0100
@@ -3,4 +3,6 @@
 ISABELLE_SLEDGEHAMMER_MASH="$COMPONENT"
 
 # MASH=yes
-MASH_PORT=9255
+if [ -z "$MASH_PORT" ]; then
+  MASH_PORT=9255
+fi
--- a/src/HOL/Tools/Sledgehammer/MaSh/src/dictionaries.py	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/MaSh/src/dictionaries.py	Mon Nov 11 17:44:21 2013 +0100
@@ -4,11 +4,12 @@
 #
 # Persistent dictionaries: accessibility, dependencies, and features.
 
-import logging,sys
+import sys
 from os.path import join
 from Queue import Queue
 from readData import create_accessible_dict,create_dependencies_dict
 from cPickle import load,dump
+from exceptions import LookupError
 
 class Dictionaries(object):
     '''
@@ -56,7 +57,6 @@
         self.changed = True
 
     def create_feature_dict(self,inputFile):
-        logger = logging.getLogger('create_feature_dict')
         self.featureDict = {}
         IS = open(inputFile,'r')
         for line in IS:
@@ -64,7 +64,7 @@
             name = line[0]
             # Name Id
             if self.nameIdDict.has_key(name):
-                logger.warning('%s appears twice in the feature file. Aborting.',name)
+                raise LookupError('%s appears twice in the feature file. Aborting.'% name)
                 sys.exit(-1)
             else:
                 self.nameIdDict[name] = self.maxNameId
@@ -134,6 +134,13 @@
                         unexpandedQueue.put(a)
         return list(accessibles)
 
+    def parse_unExpAcc(self,line):
+        try:
+            unExpAcc = [self.nameIdDict[a.strip()] for a in line.split()]            
+        except:
+            raise LookupError('Cannot find the accessibles:%s. Accessibles need to be introduced before referring to them.' % line)
+        return unExpAcc
+
     def parse_fact(self,line):
         """
         Parses a single line, extracting accessibles, features, and dependencies.
@@ -147,8 +154,9 @@
         nameId = self.get_name_id(name)
         line = line[1].split(';')
         # Accessible Ids
-        unExpAcc = [self.nameIdDict[a.strip()] for a in line[0].split()]
-        self.accessibleDict[nameId] = unExpAcc
+        #unExpAcc = [self.nameIdDict[a.strip()] for a in line[0].split()]
+        #self.accessibleDict[nameId] = unExpAcc
+        self.accessibleDict[nameId] = self.parse_unExpAcc(line[0])
         features = self.get_features(line)
         self.featureDict[nameId] = features
         self.dependenciesDict[nameId] = [self.nameIdDict[d.strip()] for d in line[2].split()]        
@@ -180,7 +188,7 @@
         name = None
         numberOfPredictions = None
 
-        # Check whether there is a problem name:
+        # How many predictions should be returned:
         tmp = line.split('#')
         if len(tmp) == 2:
             numberOfPredictions = int(tmp[0].strip())
@@ -194,8 +202,11 @@
 
         # line = accessibles;features
         line = line.split(';')
+        features = self.get_features(line)
+        
         # Accessible Ids, expand and store the accessibles.
-        unExpAcc = [self.nameIdDict[a.strip()] for a in line[0].split()]
+        #unExpAcc = [self.nameIdDict[a.strip()] for a in line[0].split()]
+        unExpAcc = self.parse_unExpAcc(line[0])        
         if len(self.expandedAccessibles.keys())>=100:
             self.expandedAccessibles = {}
             self.changed = True
@@ -205,7 +216,7 @@
                 self.expandedAccessibles[accId] = self.expand_accessibles(accIdAcc)
                 self.changed = True
         accessibles = self.expand_accessibles(unExpAcc)
-        features = self.get_features(line)
+        
         # Get hints:
         if len(line) == 3:
             hints = [self.nameIdDict[d.strip()] for d in line[2].split()]
--- a/src/HOL/Tools/Sledgehammer/MaSh/src/mash.py	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/MaSh/src/mash.py	Mon Nov 11 17:44:21 2013 +0100
@@ -22,13 +22,25 @@
 from parameters import init_parser
 
 def communicate(data,host,port):
-    sock = socket.socket(socket.AF_INET, socket.SOCK_STREAM)
+    logger = logging.getLogger('communicate')
+    sock = socket.socket(socket.AF_INET, socket.SOCK_STREAM)    
     try:
         sock.connect((host,port))
-        sock.sendall(data)
-        received = sock.recv(4194304)
-    except:
-        logger = logging.getLogger('communicate')
+        sock.sendall(data+'\n')        
+        received = ''
+        cont = True
+        counter = 0
+        while cont and counter < 100000:
+            rec = sock.recv(4096)
+            if rec.endswith('stop'):
+                cont = False
+                received += rec[:-4]
+            else:
+                received += rec
+            counter += 1
+        if rec == '':
+            logger.warning('No response from server. Check server log for details.')
+    except:        
         logger.warning('Communication with server failed.')
         received = -1
     finally:
--- a/src/HOL/Tools/Sledgehammer/MaSh/src/server.py	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/MaSh/src/server.py	Mon Nov 11 17:44:21 2013 +0100
@@ -19,7 +19,8 @@
 from stats import Statistics
 
 
-class ThreadingTCPServer(SocketServer.ThreadingTCPServer): 
+class ThreadingTCPServer(SocketServer.ThreadingTCPServer):
+    
     def __init__(self, *args, **kwargs):
         SocketServer.ThreadingTCPServer.__init__(self,*args, **kwargs)
         self.manager = Manager()
@@ -27,8 +28,17 @@
         self.idle_timeout = 28800.0 # 8 hours in seconds
         self.idle_timer = Timer(self.idle_timeout, self.shutdown)
         self.idle_timer.start()        
+        self.model = None
+        self.dicts = None
+        self.callCounter = 0
         
     def save(self):
+        if self.model == None or self.dicts == None:
+            try:
+                self.logger.warning('Cannot save nonexisting models.')
+            except:
+                pass
+            return
         # Save Models
         self.model.save(self.args.modelFile)
         self.dicts.save(self.args.dictsFile)
@@ -40,7 +50,7 @@
         self.save()          
         self.shutdown()
 
-class MaShHandler(SocketServer.BaseRequestHandler):
+class MaShHandler(SocketServer.StreamRequestHandler):
 
     def init(self,argv):
         if argv == '':
@@ -140,7 +150,7 @@
         #predictionValues = [str(x) for x in predictionValues[:numberOfPredictions]]
         #predictionsStringList = ['%s=%s' % (predictionNames[i],predictionValues[i]) for i in range(len(predictionNames))]
         #predictionsString = string.join(predictionsStringList,' ')
-        predictionsString = string.join(predictionNames,' ')
+        predictionsString = string.join(predictionNames,' ')        
         outString = '%s: %s' % (name,predictionsString)
         self.request.sendall(outString)
     
@@ -154,15 +164,15 @@
 
     def handle(self):
         # self.request is the TCP socket connected to the client
-        self.data = self.request.recv(4194304).strip()
         self.server.lock.acquire()
+        self.data = self.rfile.readline().strip()
         try:
             # Update idle shutdown timer
             self.server.idle_timer.cancel()
             self.server.idle_timer = Timer(self.server.idle_timeout, self.server.save_and_shutdown)
             self.server.idle_timer.start()        
 
-            self.startTime = time()  
+            self.startTime = time()
             if self.data == 'shutdown':
                 self.shutdown()         
             elif self.data == 'save':
@@ -189,12 +199,19 @@
             else:
                 self.request.sendall('Unspecified input format: \n%s',self.data)
             self.server.callCounter += 1
+            self.request.sendall('stop')                       
+        except: # catch exceptions
+            #print 'Caught an error. Check %s for more details' % (self.server.args.log+'server')
+            logging.exception('')
         finally:
             self.server.lock.release()
 
 if __name__ == "__main__":
-    HOST, PORT = sys.argv[1:]    
-    #HOST, PORT = "localhost", 9255
+    if not len(sys.argv[1:]) == 2:
+        print 'No Arguments for HOST and PORT found. Using localhost and 9255'
+        HOST, PORT = "localhost", 9255
+    else:
+        HOST, PORT = sys.argv[1:]
     SocketServer.TCPServer.allow_reuse_address = True
     server = ThreadingTCPServer((HOST, int(PORT)), MaShHandler)
     server.serve_forever()        
--- a/src/HOL/Tools/Sledgehammer/MaSh/src/sparseNaiveBayes.py	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/MaSh/src/sparseNaiveBayes.py	Mon Nov 11 17:44:21 2013 +0100
@@ -78,6 +78,8 @@
             self.counts[dep][0] -= 1
             for f,_w in features.items():
                 self.counts[dep][1][f] -= 1
+                if self.counts[dep][1][f] == 0:
+                    del self.counts[dep][1][f]
 
 
     def overwrite(self,problemId,newDependencies,dicts):
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -39,6 +39,7 @@
   val nearly_all_facts :
     Proof.context -> bool -> fact_override -> unit Symtab.table
     -> status Termtab.table -> thm list -> term list -> term -> raw_fact list
+  val drop_duplicate_facts : raw_fact list -> raw_fact list
 end;
 
 structure Sledgehammer_Fact : SLEDGEHAMMER_FACT =
@@ -57,6 +58,11 @@
    del : (Facts.ref * Attrib.src list) list,
    only : bool}
 
+(* gracefully handle huge background theories *)
+val max_facts_for_duplicates = 50000
+val max_facts_for_complex_check = 25000
+val max_simps_for_clasimpset = 10000
+
 (* experimental feature *)
 val instantiate_inducts =
   Attrib.setup_config_bool @{binding sledgehammer_instantiate_inducts} (K false)
@@ -102,6 +108,7 @@
                      body_type T = @{typ bool}
                    | _ => false)
 
+(* TODO: get rid of *)
 fun normalize_vars t =
   let
     fun normT (Type (s, Ts)) = fold_map normT Ts #>> curry Type s
@@ -126,25 +133,31 @@
   in fst (norm t (([], 0), ([], 0))) end
 
 fun status_of_thm css name th =
-  let val t = normalize_vars (prop_of th) in
-    (* FIXME: use structured name *)
-    if String.isSubstring ".induct" name andalso may_be_induction t then
-      Induction
-    else case Termtab.lookup css t of
-      SOME status => status
-    | NONE =>
-      let val concl = Logic.strip_imp_concl t in
-        case try (HOLogic.dest_eq o HOLogic.dest_Trueprop) concl of
-          SOME lrhss =>
-          let
-            val prems = Logic.strip_imp_prems t
-            val t' = Logic.list_implies (prems, Logic.mk_equals lrhss)
-          in
-            Termtab.lookup css t' |> the_default General
-          end
-        | NONE => General
-      end
-  end
+  if Termtab.is_empty css then
+    General
+  else
+    let val t = prop_of th in
+      (* FIXME: use structured name *)
+      if String.isSubstring ".induct" name andalso may_be_induction t then
+        Induction
+      else
+        let val t = normalize_vars t in
+          case Termtab.lookup css t of
+            SOME status => status
+          | NONE =>
+            let val concl = Logic.strip_imp_concl t in
+              case try (HOLogic.dest_eq o HOLogic.dest_Trueprop) concl of
+                SOME lrhss =>
+                let
+                  val prems = Logic.strip_imp_prems t
+                  val t' = Logic.list_implies (prems, Logic.mk_equals lrhss)
+                in
+                  Termtab.lookup css t' |> the_default General
+                end
+              | NONE => General
+            end
+        end
+    end
 
 fun stature_of_thm global assms chained css name th =
   (scope_of_thm global assms chained th, status_of_thm css name th)
@@ -219,10 +232,11 @@
 
 val sep_that = Long_Name.separator ^ Obtain.thatN
 
+val skolem_thesis = Name.skolem Auto_Bind.thesisN
+
 fun is_that_fact th =
-  String.isSuffix sep_that (Thm.get_name_hint th)
-  andalso exists_subterm (fn Free (s, _) => s = Name.skolem Auto_Bind.thesisN
-                           | _ => false) (prop_of th)
+  exists_subterm (fn Free (s, _) => s = skolem_thesis | _ => false) (prop_of th)
+  andalso String.isSuffix sep_that (Thm.get_name_hint th)
 
 datatype interest = Deal_Breaker | Interesting | Boring
 
@@ -295,9 +309,7 @@
 fun backquote_term ctxt = close_form #> hackish_string_of_term ctxt #> backquote
 fun backquote_thm ctxt = backquote_term ctxt o prop_of
 
-(* gracefully handle huge background theories *)
-val max_simps_for_clasimpset = 10000
-
+(* TODO: rewrite to use nets and/or to reuse existing data structures *)
 fun clasimpset_rule_table_of ctxt =
   let val simps = ctxt |> simpset_of |> dest_ss |> #simps in
     if length simps >= max_simps_for_clasimpset then
@@ -336,18 +348,16 @@
       end
   end
 
-fun normalize_eq (t as @{const Trueprop}
-        $ ((t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2)) =
-    if Term_Ord.fast_term_ord (t1, t2) <> GREATER then t
-    else HOLogic.mk_Trueprop (t0 $ t2 $ t1)
-  | normalize_eq (t as @{const Trueprop} $ (@{const Not}
+fun normalize_eq (@{const Trueprop} $ (t as (t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2)) =
+    if Term_Ord.fast_term_ord (t1, t2) <> GREATER then t else t0 $ t2 $ t1
+  | normalize_eq (@{const Trueprop} $ (t as @{const Not}
         $ ((t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2))) =
-    if Term_Ord.fast_term_ord (t1, t2) <> GREATER then t
-    else HOLogic.mk_Trueprop (HOLogic.mk_not (t0 $ t2 $ t1))
+    if Term_Ord.fast_term_ord (t1, t2) <> GREATER then t else HOLogic.mk_not (t0 $ t2 $ t1)
+  | normalize_eq (Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2) =
+    (if Term_Ord.fast_term_ord (t1, t2) <> GREATER then (t1, t2) else (t2, t1))
+    |> (fn (t1, t2) => HOLogic.eq_const T $ t1 $ t2)
   | normalize_eq t = t
 
-val normalize_eq_vars = normalize_eq #> normalize_vars
-
 fun if_thm_before th th' =
   if Theory.subthy (pairself Thm.theory_of_thm (th, th')) then th else th'
 
@@ -361,30 +371,22 @@
 
 fun build_name_tables name_of facts =
   let
-    fun cons_thm (_, th) =
-      Termtab.cons_list (normalize_eq_vars (prop_of th), th)
+    fun cons_thm (_, th) = Termtab.cons_list (normalize_vars (normalize_eq (prop_of th)), th)
     fun add_plain canon alias =
-      Symtab.update (Thm.get_name_hint alias,
-                     name_of (if_thm_before canon alias))
+      Symtab.update (Thm.get_name_hint alias, name_of (if_thm_before canon alias))
     fun add_plains (_, aliases as canon :: _) = fold (add_plain canon) aliases
-    fun add_inclass (name, target) =
-      fold (fn s => Symtab.update (s, target)) (un_class_ify name)
+    fun add_inclass (name, target) = fold (fn s => Symtab.update (s, target)) (un_class_ify name)
     val prop_tab = fold cons_thm facts Termtab.empty
     val plain_name_tab = Termtab.fold add_plains prop_tab Symtab.empty
     val inclass_name_tab = Symtab.fold add_inclass plain_name_tab Symtab.empty
   in (plain_name_tab, inclass_name_tab) end
 
-fun keyed_distinct key_of xs =
-  let val tab = fold (Termtab.default o `key_of) xs Termtab.empty in
-    Termtab.fold (cons o snd) tab []
-  end
-
-fun hashed_keyed_distinct hash_of key_of xs =
-  let
-    val ys = map (`hash_of) xs
-    val sorted_ys = sort (int_ord o pairself fst) ys
-    val grouped_ys = AList.coalesce (op =) sorted_ys
-  in maps (keyed_distinct key_of o snd) grouped_ys end
+fun fact_distinct eq facts =
+  fold (fn fact as (_, th) =>
+      Net.insert_term_safe (eq o pairself (normalize_eq o prop_of o snd))
+        (normalize_eq (prop_of th), fact))
+    facts Net.empty
+  |> Net.entries
 
 fun struct_induct_rule_on th =
   case Logic.strip_horn (prop_of th) of
@@ -447,9 +449,6 @@
 
 fun fact_count facts = Facts.fold_static (K (Integer.add 1)) facts 0
 
-(* gracefully handle huge background theories *)
-val max_facts_for_complex_check = 25000
-
 fun all_facts ctxt generous ho_atp reserved add_ths chained css =
   let
     val thy = Proof_Context.theory_of ctxt
@@ -473,12 +472,12 @@
       Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts)
     val is_blacklisted_or_something = is_blacklisted_or_something ctxt ho_atp
     fun add_facts global foldx facts =
-      foldx (fn (name0, ths) =>
+      foldx (fn (name0, ths) => fn accum =>
         if name0 <> "" andalso
            forall (not o member Thm.eq_thm_prop add_ths) ths andalso
            (Facts.is_concealed facts name0 orelse
             (not generous andalso is_blacklisted_or_something name0)) then
-          I
+          accum
         else
           let
             val n = length ths
@@ -488,32 +487,30 @@
                 NONE => false
               | SOME ths' => eq_list Thm.eq_thm_prop (ths, ths')
           in
-            pair n
-            #> fold_rev (fn th => fn (j, accum) =>
-                   (j - 1,
-                    if not (member Thm.eq_thm_prop add_ths th) andalso
-                       (is_likely_tautology_too_meta_or_too_technical th orelse
-                        is_too_complex (prop_of th)) then
-                      accum
-                    else
-                      let
-                        val new =
-                          (((fn () =>
-                                if name0 = "" then
-                                  backquote_thm ctxt th
-                                else
-                                  [Facts.extern ctxt facts name0,
-                                   Name_Space.extern ctxt full_space name0]
-                                  |> distinct (op =)
-                                  |> find_first check_thms
-                                  |> the_default name0
-                                  |> make_name reserved multi j),
-                             stature_of_thm global assms chained css name0 th),
-                           th)
-                      in
-                        accum |> (if multi then apsnd else apfst) (cons new)
-                      end)) ths
-            #> snd
+            snd (fold_rev (fn th => fn (j, accum as (uni_accum, multi_accum)) =>
+              (j - 1,
+               if not (member Thm.eq_thm_prop add_ths th) andalso
+                  (is_likely_tautology_too_meta_or_too_technical th orelse
+                   is_too_complex (prop_of th)) then
+                 accum
+               else
+                 let
+                   val new =
+                     (((fn () =>
+                           if name0 = "" then
+                             backquote_thm ctxt th
+                           else
+                             [Facts.extern ctxt facts name0,
+                              Name_Space.extern ctxt full_space name0]
+                             |> find_first check_thms
+                             |> the_default name0
+                             |> make_name reserved multi j),
+                        stature_of_thm global assms chained css name0 th),
+                      th)
+                 in
+                   if multi then (uni_accum, new :: multi_accum)
+                   else (new :: uni_accum, multi_accum)
+                 end)) ths (n, accum))
           end)
   in
     (* The single-theorem names go before the multiple-theorem ones (e.g.,
@@ -530,6 +527,7 @@
     []
   else
     let
+      val thy = Proof_Context.theory_of ctxt
       val chained =
         chained
         |> maps (fn th => insert Thm.eq_thm_prop (zero_var_indexes th) [th])
@@ -538,14 +536,25 @@
          maps (map (fn ((name, stature), th) => ((K name, stature), th))
                o fact_of_ref ctxt reserved chained css) add
        else
-         let val (add, del) = pairself (Attrib.eval_thms ctxt) (add, del) in
-           all_facts ctxt false ho_atp reserved add chained css
-           |> filter_out
-                  ((member Thm.eq_thm_prop del orf No_ATPs.member ctxt) o snd)
-           |> hashed_keyed_distinct (size_of_term o prop_of o snd)
-                  (normalize_eq_vars o prop_of o snd)
+        (* The "fact_distinct" call would have cleaner semantics if it called "Pattern.equiv"
+           instead of "Pattern.matches", but it would also be slower and less precise.
+           "Pattern.matches" throws out theorems that are strict instances of other theorems, but
+           only if the instance is met after the general version. *)
+         let
+           val (add, del) = pairself (Attrib.eval_thms ctxt) (add, del)
+           val facts =
+             all_facts ctxt false ho_atp reserved add chained css
+             |> filter_out ((member Thm.eq_thm_prop del orf
+               (No_ATPs.member ctxt andf not o member Thm.eq_thm_prop add)) o snd)
+         in
+           facts
          end)
       |> maybe_instantiate_inducts ctxt hyp_ts concl_t
     end
 
+fun drop_duplicate_facts facts =
+  let val num_facts = length facts in
+    facts |> num_facts <= max_facts_for_duplicates ? fact_distinct (op aconv)
+  end
+
 end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -183,6 +183,7 @@
    read correctly. *)
 val implode_param = strip_spaces_except_between_idents o space_implode " "
 
+(* FIXME: use "Generic_Data" *)
 structure Data = Theory_Data
 (
   type T = raw_param list
@@ -196,22 +197,24 @@
 
 (* The first ATP of the list is used by Auto Sledgehammer. Because of the low
    timeout, it makes sense to put E first. *)
-fun default_provers_param_value ctxt =
-  [eN, spassN, vampireN, z3N, e_sineN, yicesN]
+fun default_provers_param_value mode ctxt =
+  [eN, spassN, z3N, vampireN, e_sineN, yicesN]
   |> map_filter (remotify_prover_if_not_installed ctxt)
-  |> take (Multithreading.max_threads_value ())
+  (* In "try" mode, leave at least one thread to another slow tool (e.g. Nitpick) *)
+  |> take (Multithreading.max_threads_value () - (if mode = Try then 1 else 0))
   |> implode_param
 
 fun set_default_raw_param param thy =
   let val ctxt = Proof_Context.init_global thy in
     thy |> Data.map (AList.update (op =) (normalize_raw_param ctxt param))
   end
-fun default_raw_params ctxt =
+
+fun default_raw_params mode ctxt =
   let val thy = Proof_Context.theory_of ctxt in
     Data.get thy
     |> fold (AList.default (op =))
             [("provers", [case !provers of
-                            "" => default_provers_param_value ctxt
+                            "" => default_provers_param_value mode ctxt
                           | s => s]),
              ("timeout", let val timeout = Options.default_int @{option sledgehammer_timeout} in
                            [if timeout <= 0 then "none"
@@ -278,7 +281,9 @@
     val lam_trans = lookup_option lookup_string "lam_trans"
     val uncurried_aliases = lookup_option lookup_bool "uncurried_aliases"
     val learn = lookup_bool "learn"
-    val fact_filter = lookup_option lookup_string "fact_filter"
+    val fact_filter =
+      lookup_option lookup_string "fact_filter"
+      |> mode = Auto_Try ? (fn NONE => SOME mepoN | sf => sf)
     val max_facts = lookup_option lookup_int "max_facts"
     val fact_thresholds = lookup_real_pair "fact_thresholds"
     val max_mono_iters = lookup_option lookup_int "max_mono_iters"
@@ -288,12 +293,10 @@
     val isar_compress = Real.max (1.0, lookup_real "isar_compress")
     val isar_try0 = lookup_bool "isar_try0"
     val slice = mode <> Auto_Try andalso lookup_bool "slice"
-    val minimize =
-      if mode = Auto_Try then NONE else lookup_option lookup_bool "minimize"
+    val minimize = if mode = Auto_Try then NONE else lookup_option lookup_bool "minimize"
     val timeout = if mode = Auto_Try then NONE else lookup_time "timeout"
     val preplay_timeout =
-      if mode = Auto_Try then SOME Time.zeroTime
-      else lookup_time "preplay_timeout"
+      if mode = Auto_Try then SOME Time.zeroTime else lookup_time "preplay_timeout"
     val expect = lookup_string "expect"
   in
     {debug = debug, verbose = verbose, overlord = overlord, spy = spy, blocking = blocking,
@@ -305,7 +308,7 @@
      timeout = timeout, preplay_timeout = preplay_timeout, expect = expect}
   end
 
-fun get_params mode = extract_params mode o default_raw_params
+fun get_params mode = extract_params mode o default_raw_params mode
 fun default_params thy = get_params Normal thy o map (apsnd single)
 
 (* Sledgehammer the given subgoal *)
@@ -385,6 +388,7 @@
        else
          ();
        mash_learn ctxt
+           (* TODO: Use MaSh mode instead and have the special defaults hardcoded in "get_params" *)
            (get_params Normal ctxt
                 ([("timeout",
                    [string_of_real default_learn_prover_timeout]),
@@ -415,7 +419,7 @@
        #> tap (fn thy =>
                   let val ctxt = Proof_Context.init_global thy in
                     writeln ("Default parameters for Sledgehammer:\n" ^
-                             (case default_raw_params ctxt |> rev of
+                             (case rev (default_raw_params Normal ctxt) of
                                 [] => "none"
                               | params =>
                                 params |> map string_of_raw_param
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -11,7 +11,6 @@
   type fact = Sledgehammer_Fact.fact
   type fact_override = Sledgehammer_Fact.fact_override
   type params = Sledgehammer_Provers.params
-  type relevance_fudge = Sledgehammer_Provers.relevance_fudge
   type prover_result = Sledgehammer_Provers.prover_result
 
   val trace : bool Config.T
@@ -61,10 +60,10 @@
   val thm_less : thm * thm -> bool
   val goal_of_thm : theory -> thm -> thm
   val run_prover_for_mash :
-    Proof.context -> params -> string -> fact list -> thm -> prover_result
+    Proof.context -> params -> string -> string -> fact list -> thm -> prover_result
   val features_of :
-    Proof.context -> string -> theory -> int -> int Symtab.table -> stature
-    -> term list -> (string * real) list
+    Proof.context -> theory -> int -> int Symtab.table -> stature -> term list ->
+    (string * real) list
   val trim_dependencies : string list -> string list option
   val isar_dependencies_of :
     string Symtab.table * string Symtab.table -> thm -> string list
@@ -180,11 +179,10 @@
       (if background then " &" else "")
     fun run_on () =
       (Isabelle_System.bash command
-       |> tap (fn _ => trace_msg ctxt (fn () =>
-              case try File.read (Path.explode err_file) of
-                NONE => "Done"
-              | SOME "" => "Done"
-              | SOME s => "Error: " ^ elide_string 1000 s));
+       |> tap (fn _ =>
+            case try File.read (Path.explode err_file) |> the_default "" of
+              "" => trace_msg ctxt (K "Done")
+            | s => warning ("MaSh error: " ^ elide_string 1000 s));
        read_suggs (fn () => try File.read_lines sugg_path |> these))
     fun clean_up () =
       if overlord then ()
@@ -514,20 +512,14 @@
       val mess = mess |> map (apsnd (apfst (normalize_scores max_facts)))
       fun score_in fact (global_weight, (sels, unks)) =
         let
-          fun score_at j =
-            case try (nth sels) j of
-              SOME (_, score) => SOME (global_weight * score)
-            | NONE => NONE
+          val score_at = try (nth sels) #> Option.map (fn (_, score) => global_weight * score)
         in
           case find_index (curry fact_eq fact o fst) sels of
-            ~1 => (case find_index (curry fact_eq fact) unks of
-                     ~1 => SOME 0.0
-                   | _ => NONE)
+            ~1 => if member fact_eq unks fact then NONE else SOME 0.0
           | rank => score_at rank
         end
       fun weight_of fact = mess |> map_filter (score_in fact) |> scaled_avg
-      val facts =
-        fold (union fact_eq o map fst o take max_facts o fst o snd) mess []
+      val facts = fold (union fact_eq o map fst o take max_facts o fst o snd) mess []
     in
       facts |> map (`weight_of) |> sort (int_ord o swap o pairself fst)
             |> map snd |> take max_facts
@@ -573,11 +565,11 @@
 
 fun goal_of_thm thy = prop_of #> freeze #> cterm_of thy #> Goal.init
 
-fun run_prover_for_mash ctxt params prover facts goal =
+fun run_prover_for_mash ctxt params prover goal_name facts goal =
   let
     val problem =
-      {state = Proof.init ctxt, goal = goal, subgoal = 1, subgoal_count = 1,
-       factss = [("", facts)]}
+      {comment = "Goal: " ^ goal_name, state = Proof.init ctxt, goal = goal, subgoal = 1,
+       subgoal_count = 1, factss = [("", facts)]}
   in
     get_minimizing_prover ctxt MaSh (K (K ())) prover params (K (K (K "")))
                           problem
@@ -585,9 +577,6 @@
 
 val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}]
 
-val logical_consts =
-  [@{const_name prop}, @{const_name Pure.conjunction}] @ atp_logical_consts
-
 val pat_tvar_prefix = "_"
 val pat_var_prefix = "_"
 
@@ -608,15 +597,10 @@
 
 val max_pat_breadth = 10 (* FUDGE *)
 
-fun term_features_of ctxt prover thy_name num_facts const_tab term_max_depth
-                     type_max_depth ts =
+fun term_features_of ctxt thy_name num_facts const_tab term_max_depth type_max_depth ts =
   let
     val thy = Proof_Context.theory_of ctxt
 
-    fun is_built_in (x as (s, _)) args =
-      if member (op =) logical_consts s then (true, args)
-      else is_built_in_const_of_prover ctxt prover x args
-
     val fixes = map snd (Variable.dest_fixes ctxt)
     val classes = Sign.classes_of thy
 
@@ -660,11 +644,10 @@
          let val count = Symtab.lookup const_tab s |> the_default 1 in
            Real.fromInt num_facts / Real.fromInt count (* FUDGE *)
          end)
-    fun pattify_term _ _ 0 _ = []
-      | pattify_term _ args _ (Const (x as (s, _))) =
-        if fst (is_built_in x args) then []
-        else [(massage_long_name s, weight_of_const s)]
-      | pattify_term _ _ _ (Free (s, T)) =
+    fun pattify_term _ 0 _ = []
+      | pattify_term _ _ (Const (s, _)) =
+        if is_widely_irrelevant_const s then [] else [(massage_long_name s, weight_of_const s)]
+      | pattify_term _ _ (Free (s, T)) =
         maybe_singleton_str pat_var_prefix (crude_str_of_typ T)
         |> map (rpair 0.0)
         |> (if member (op =) fixes s then
@@ -672,36 +655,31 @@
                   (thy_name ^ Long_Name.separator ^ s)))
             else
               I)
-      | pattify_term _ _ _ (Var (_, T)) =
-        maybe_singleton_str pat_var_prefix (crude_str_of_typ T)
-        |> map (rpair 0.0)
-      | pattify_term Ts _ _ (Bound j) =
-        maybe_singleton_str pat_var_prefix (crude_str_of_typ (nth Ts j))
-        |> map (rpair 0.0)
-      | pattify_term Ts args depth (t $ u) =
+      | pattify_term _ _ (Var (_, T)) =
+        maybe_singleton_str pat_var_prefix (crude_str_of_typ T) |> map (rpair 0.0)
+      | pattify_term Ts _ (Bound j) =
+        maybe_singleton_str pat_var_prefix (crude_str_of_typ (nth Ts j)) |> map (rpair 0.0)
+      | pattify_term Ts depth (t $ u) =
         let
-          val ps = take max_pat_breadth (pattify_term Ts (u :: args) depth t)
-          val qs =
-            take max_pat_breadth (("", 0.0) :: pattify_term Ts [] (depth - 1) u)
+          val ps = take max_pat_breadth (pattify_term Ts depth t)
+          val qs = take max_pat_breadth (("", 0.0) :: pattify_term Ts (depth - 1) u)
         in
           map_product (fn ppw as (p, pw) =>
               fn ("", _) => ppw
                | (q, qw) => (p ^ "(" ^ q ^ ")", pw + qw)) ps qs
         end
-      | pattify_term _ _ _ _ = []
-    fun add_term_pat Ts = union (eq_fst (op =)) oo pattify_term Ts []
+      | pattify_term _ _ _ = []
+    fun add_term_pat Ts = union (eq_fst (op =)) oo pattify_term Ts
     fun add_term_pats _ 0 _ = I
       | add_term_pats Ts depth t =
         add_term_pat Ts depth t #> add_term_pats Ts (depth - 1) t
     fun add_term Ts = add_term_pats Ts term_max_depth
     fun add_subterms Ts t =
       case strip_comb t of
-        (Const (x as (_, T)), args) =>
-        let val (built_in, args) = is_built_in x args in
-          (not built_in ? add_term Ts t)
-          #> add_subtypes T
-          #> fold (add_subterms Ts) args
-        end
+        (Const (s, T), args) =>
+        (not (is_widely_irrelevant_const s) ? add_term Ts t)
+        #> add_subtypes T
+        #> fold (add_subterms Ts) args
       | (head, args) =>
         (case head of
            Free (_, T) => add_term Ts t #> add_subtypes T
@@ -715,11 +693,10 @@
 val type_max_depth = 1
 
 (* TODO: Generate type classes for types? *)
-fun features_of ctxt prover thy num_facts const_tab (scope, _) ts =
+fun features_of ctxt thy num_facts const_tab (scope, _) ts =
   let val thy_name = Context.theory_name thy in
     thy_feature_of thy_name ::
-    term_features_of ctxt prover thy_name num_facts const_tab term_max_depth
-                     type_max_depth ts
+    term_features_of ctxt thy_name num_facts const_tab term_max_depth type_max_depth ts
     |> scope <> Global ? cons local_feature
   end
 
@@ -727,7 +704,7 @@
    isn't much to learn from such proofs. *)
 val max_dependencies = 20
 
-val prover_default_max_facts = 50
+val prover_default_max_facts = 25
 
 (* "type_definition_xxx" facts are characterized by their use of "CollectI". *)
 val typedef_dep = nickname_of_thm @{thm CollectI}
@@ -785,6 +762,7 @@
     let
       val thy = Proof_Context.theory_of ctxt
       val goal = goal_of_thm thy th
+      val name = nickname_of_thm th
       val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal goal 1 ctxt
       val facts = facts |> filter (fn (_, th') => thm_less (th', th))
       fun nickify ((_, stature), th) = ((nickname_of_thm th, stature), th)
@@ -795,24 +773,23 @@
         else case find_first (is_dep dep) facts of
           SOME ((_, status), th) => accum @ [(("", status), th)]
         | NONE => accum (* shouldn't happen *)
-      val facts =
+      val mepo_facts =
         facts
-        |> mepo_suggested_facts ctxt params prover
-               (max_facts |> the_default prover_default_max_facts) NONE hyp_ts
-               concl_t
+        |> mepo_suggested_facts ctxt params (max_facts |> the_default prover_default_max_facts) NONE
+             hyp_ts concl_t
+      val facts =
+        mepo_facts
         |> fold (add_isar_dep facts) isar_deps
         |> map nickify
+      val num_isar_deps = length isar_deps
     in
       if verbose andalso auto_level = 0 then
-        let val num_facts = length facts in
-          "MaSh: " ^ quote prover ^ " on " ^ quote (nickname_of_thm th) ^
-          " with " ^ string_of_int num_facts ^ " fact" ^ plural_s num_facts ^
-          "."
-          |> Output.urgent_message
-        end
+        "MaSh: " ^ quote prover ^ " on " ^ quote name ^ " with " ^ string_of_int num_isar_deps ^
+        " + " ^ string_of_int (length facts - num_isar_deps) ^ " facts."
+        |> Output.urgent_message
       else
         ();
-      case run_prover_for_mash ctxt params prover facts goal of
+      case run_prover_for_mash ctxt params prover name facts goal of
         {outcome = NONE, used_facts, ...} =>
         (if verbose andalso auto_level = 0 then
            let val num_facts = length used_facts in
@@ -935,23 +912,22 @@
 
 val max_proximity_facts = 100
 
-fun find_mash_suggestions _ _ [] _ _ raw_unknown = ([], raw_unknown)
-  | find_mash_suggestions ctxt max_facts suggs facts chained raw_unknown =
-    let
-      val inter_fact = inter (eq_snd Thm.eq_thm_prop)
-      val raw_mash = find_suggested_facts ctxt facts suggs
-      val proximate = take max_proximity_facts facts
-      val unknown_chained = inter_fact raw_unknown chained
-      val unknown_proximate = inter_fact raw_unknown proximate
-      val mess =
-        [(0.9 (* FUDGE *), (map (rpair 1.0) unknown_chained, [])),
-         (0.4 (* FUDGE *), (weight_facts_smoothly unknown_proximate, [])),
-         (0.1 (* FUDGE *), (weight_facts_steeply raw_mash, raw_unknown))]
-      val unknown =
-        raw_unknown
-        |> fold (subtract (eq_snd Thm.eq_thm_prop))
-                [unknown_chained, unknown_proximate]
-    in (mesh_facts (eq_snd Thm.eq_thm_prop) max_facts mess, unknown) end
+fun find_mash_suggestions ctxt max_facts suggs facts chained raw_unknown =
+  let
+    val inter_fact = inter (eq_snd Thm.eq_thm_prop)
+    val raw_mash = find_suggested_facts ctxt facts suggs
+    val proximate = take max_proximity_facts facts
+    val unknown_chained = inter_fact raw_unknown chained
+    val unknown_proximate = inter_fact raw_unknown proximate
+    val mess =
+      [(0.9 (* FUDGE *), (map (rpair 1.0) unknown_chained, [])),
+       (0.4 (* FUDGE *), (weight_facts_smoothly unknown_proximate, [])),
+       (0.1 (* FUDGE *), (weight_facts_steeply raw_mash, raw_unknown))]
+    val unknown =
+      raw_unknown
+      |> fold (subtract (eq_snd Thm.eq_thm_prop))
+              [unknown_chained, unknown_proximate]
+  in (mesh_facts (eq_snd Thm.eq_thm_prop) max_facts mess, unknown) end
 
 fun add_const_counts t =
   fold (fn s => Symtab.map_default (s, 0) (Integer.add 1))
@@ -966,22 +942,23 @@
     val chained = facts |> filter (fn ((_, (scope, _)), _) => scope = Chained)
     val num_facts = length facts
     val const_tab = fold (add_const_counts o prop_of o snd) facts Symtab.empty
+
     fun fact_has_right_theory (_, th) =
       thy_name = Context.theory_name (theory_of_thm th)
     fun chained_or_extra_features_of factor (((_, stature), th), weight) =
       [prop_of th]
-      |> features_of ctxt prover (theory_of_thm th) num_facts const_tab stature
+      |> features_of ctxt (theory_of_thm th) num_facts const_tab stature
       |> map (apsnd (fn r => weight * factor * r))
+
     val (access_G, suggs) =
       peek_state ctxt overlord (fn {access_G, ...} =>
           if Graph.is_empty access_G then
-            (access_G, [])
+            (trace_msg ctxt (K "Nothing has been learned yet"); (access_G, []))
           else
             let
               val parents = maximal_wrt_access_graph access_G facts
               val goal_feats =
-                features_of ctxt prover thy num_facts const_tab (Local, General)
-                            (concl_t :: hyp_ts)
+                features_of ctxt thy num_facts const_tab (Local, General) (concl_t :: hyp_ts)
               val chained_feats =
                 chained
                 |> map (rpair 1.0)
@@ -1050,9 +1027,7 @@
         let
           val thy = Proof_Context.theory_of ctxt
           val name = freshish_name ()
-          val feats =
-            features_of ctxt prover thy 0 Symtab.empty (Local, General) [t]
-            |> map fst
+          val feats = features_of ctxt thy 0 Symtab.empty (Local, General) [t] |> map fst
         in
           peek_state ctxt overlord (fn {access_G, ...} =>
               let
@@ -1151,9 +1126,7 @@
             let
               val name = nickname_of_thm th
               val feats =
-                features_of ctxt prover (theory_of_thm th) 0 Symtab.empty
-                            stature [prop_of th]
-                |> map fst
+                features_of ctxt (theory_of_thm th) 0 Symtab.empty stature [prop_of th] |> map fst
               val deps = deps_of status th |> these
               val n = n |> not (null deps) ? Integer.add 1
               val learns = (name, parents, feats, deps) :: learns
@@ -1215,7 +1188,7 @@
                      Isar_Proof => 0
                    | Automatic_Proof => 2 * max_isar
                    | Isar_Proof_wegen_Prover_Flop => max_isar)
-                - 500 * length (isar_dependencies_of name_tabs th)
+                - 100 * length (isar_dependencies_of name_tabs th)
               val old_facts =
                 facts |> filter is_in_access_G
                       |> map (`priority_of)
@@ -1228,7 +1201,7 @@
       in
         if verbose orelse auto_level < 2 then
           "Learned " ^ string_of_int n ^ " nontrivial " ^
-          (if run_prover then "automatic" else "Isar") ^ " proof" ^ plural_s n ^
+          (if run_prover then "automatic and " else "") ^ "Isar proof" ^ plural_s n ^
           (if verbose then " in " ^ string_of_time (Timer.checkRealTimer timer)
            else "") ^ "."
         else
@@ -1236,15 +1209,14 @@
       end
   end
 
-fun mash_learn ctxt (params as {provers, timeout, max_facts, ...}) fact_override
-               chained run_prover =
+fun mash_learn ctxt (params as {provers, timeout, ...}) fact_override chained run_prover =
   let
     val css = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
     val ctxt = ctxt |> Config.put instantiate_inducts false
     val facts =
       nearly_all_facts ctxt false fact_override Symtab.empty css chained []
                        @{prop True}
-      |> (case max_facts of NONE => I | SOME n => take n)
+      |> sort (crude_thm_ord o pairself snd o swap)
     val num_facts = length facts
     val prover = hd provers
     fun learn auto_level run_prover =
@@ -1345,42 +1317,36 @@
              if mash_can_suggest_facts ctxt overlord then meshN else mepoN)
           else
             (false, mepoN)
+
+      val unique_facts = drop_duplicate_facts facts
       val add_ths = Attrib.eval_thms ctxt add
+
       fun in_add (_, th) = member Thm.eq_thm_prop add_ths th
       fun add_and_take accepts =
         (case add_ths of
            [] => accepts
-         | _ => (facts |> filter in_add |> map fact_of_raw_fact) @
+         | _ => (unique_facts |> filter in_add |> map fact_of_raw_fact) @
                 (accepts |> filter_out in_add))
         |> take max_facts
       fun mepo () =
-        mepo_suggested_facts ctxt params prover max_facts NONE hyp_ts concl_t
-                             facts
-        |> weight_facts_steeply
+        (mepo_suggested_facts ctxt params max_facts NONE hyp_ts concl_t unique_facts
+         |> weight_facts_steeply, [])
       fun mash () =
-        mash_suggested_facts ctxt params prover (generous_max_facts max_facts)
-            hyp_ts concl_t facts
+        mash_suggested_facts ctxt params prover (generous_max_facts max_facts) hyp_ts concl_t facts
         |>> weight_facts_steeply
       val mess =
         (* the order is important for the "case" expression below *)
-        [] |> (if effective_fact_filter <> mepoN then
-                 cons (mash_weight, (mash ()))
-               else
-                 I)
-           |> (if effective_fact_filter <> mashN then
-                 cons (mepo_weight, (mepo (), []))
-               else
-                 I)
-      val mesh =
-        mesh_facts (eq_snd Thm.eq_thm_prop) max_facts mess
-        |> add_and_take
+        [] |> effective_fact_filter <> mepoN ? cons (mash_weight, mash)
+           |> effective_fact_filter <> mashN ? cons (mepo_weight, mepo)
+           |> Par_List.map (apsnd (fn f => f ()))
+      val mesh = mesh_facts (eq_snd Thm.eq_thm_prop) max_facts mess |> add_and_take
     in
       if save then MaSh.save ctxt overlord else ();
       case (fact_filter, mess) of
         (NONE, [(_, (mepo, _)), (_, (mash, _))]) =>
         [(meshN, mesh), (mepoN, mepo |> map fst |> add_and_take),
          (mashN, mash |> map fst |> add_and_take)]
-      | _ => [("", mesh)]
+      | _ => [(effective_fact_filter, mesh)]
     end
 
 fun kill_learners ctxt ({overlord, ...} : params) =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_mepo.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_mepo.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -11,13 +11,33 @@
   type raw_fact = Sledgehammer_Fact.raw_fact
   type fact = Sledgehammer_Fact.fact
   type params = Sledgehammer_Provers.params
-  type relevance_fudge = Sledgehammer_Provers.relevance_fudge
+
+  type relevance_fudge =
+    {local_const_multiplier : real,
+     worse_irrel_freq : real,
+     higher_order_irrel_weight : real,
+     abs_rel_weight : real,
+     abs_irrel_weight : real,
+     theory_const_rel_weight : real,
+     theory_const_irrel_weight : real,
+     chained_const_irrel_weight : real,
+     intro_bonus : real,
+     elim_bonus : real,
+     simp_bonus : real,
+     local_bonus : real,
+     assum_bonus : real,
+     chained_bonus : real,
+     max_imperfect : real,
+     max_imperfect_exp : real,
+     threshold_divisor : real,
+     ridiculous_threshold : real}
 
   val trace : bool Config.T
   val pseudo_abs_name : string
+  val default_relevance_fudge : relevance_fudge
   val mepo_suggested_facts :
-    Proof.context -> params -> string -> int -> relevance_fudge option
-    -> term list -> term -> raw_fact list -> fact list
+    Proof.context -> params -> int -> relevance_fudge option -> term list -> term ->
+    raw_fact list -> fact list
 end;
 
 structure Sledgehammer_MePo : SLEDGEHAMMER_MEPO =
@@ -36,6 +56,47 @@
 val pseudo_abs_name = sledgehammer_prefix ^ "abs"
 val theory_const_suffix = Long_Name.separator ^ " 1"
 
+type relevance_fudge =
+  {local_const_multiplier : real,
+   worse_irrel_freq : real,
+   higher_order_irrel_weight : real,
+   abs_rel_weight : real,
+   abs_irrel_weight : real,
+   theory_const_rel_weight : real,
+   theory_const_irrel_weight : real,
+   chained_const_irrel_weight : real,
+   intro_bonus : real,
+   elim_bonus : real,
+   simp_bonus : real,
+   local_bonus : real,
+   assum_bonus : real,
+   chained_bonus : real,
+   max_imperfect : real,
+   max_imperfect_exp : real,
+   threshold_divisor : real,
+   ridiculous_threshold : real}
+
+(* FUDGE *)
+val default_relevance_fudge =
+  {local_const_multiplier = 1.5,
+   worse_irrel_freq = 100.0,
+   higher_order_irrel_weight = 1.05,
+   abs_rel_weight = 0.5,
+   abs_irrel_weight = 2.0,
+   theory_const_rel_weight = 0.5,
+   theory_const_irrel_weight = 0.25,
+   chained_const_irrel_weight = 0.25,
+   intro_bonus = 0.15,
+   elim_bonus = 0.15,
+   simp_bonus = 0.15,
+   local_bonus = 0.55,
+   assum_bonus = 1.05,
+   chained_bonus = 1.5,
+   max_imperfect = 11.5,
+   max_imperfect_exp = 1.0,
+   threshold_divisor = 2.0,
+   ridiculous_threshold = 0.1}
+
 fun order_of_type (Type (@{type_name fun}, [T1, T2])) =
     Int.max (order_of_type T1 + 1, order_of_type T2)
   | order_of_type (Type (_, Ts)) = fold (Integer.max o order_of_type) Ts 0
@@ -43,24 +104,26 @@
 
 (* An abstraction of Isabelle types and first-order terms *)
 datatype pattern = PVar | PApp of string * pattern list
-datatype ptype = PType of int * pattern list
+datatype ptype = PType of int * typ list
 
-fun string_of_pattern PVar = "_"
-  | string_of_pattern (PApp (s, ps)) =
-    if null ps then s else s ^ string_of_patterns ps
-and string_of_patterns ps = "(" ^ commas (map string_of_pattern ps) ^ ")"
-fun string_of_ptype (PType (_, ps)) = string_of_patterns ps
+fun string_of_patternT (TVar _) = "_"
+  | string_of_patternT (Type (s, ps)) =
+    if null ps then s else s ^ string_of_patternsT ps
+  | string_of_patternT (TFree (s, _)) = s
+and string_of_patternsT ps = "(" ^ commas (map string_of_patternT ps) ^ ")"
+fun string_of_ptype (PType (_, ps)) = string_of_patternsT ps
 
 (*Is the second type an instance of the first one?*)
-fun match_pattern (PVar, _) = true
-  | match_pattern (PApp _, PVar) = false
-  | match_pattern (PApp (s, ps), PApp (t, qs)) =
-    s = t andalso match_patterns (ps, qs)
-and match_patterns (_, []) = true
-  | match_patterns ([], _) = false
-  | match_patterns (p :: ps, q :: qs) =
-    match_pattern (p, q) andalso match_patterns (ps, qs)
-fun match_ptype (PType (_, ps), PType (_, qs)) = match_patterns (ps, qs)
+fun match_patternT (TVar _, _) = true
+  | match_patternT (Type (s, ps), Type (t, qs)) =
+    s = t andalso match_patternsT (ps, qs)
+  | match_patternT (TFree (s, _), TFree (t, _)) = s = t
+  | match_patternT (_, _) = false
+and match_patternsT (_, []) = true
+  | match_patternsT ([], _) = false
+  | match_patternsT (p :: ps, q :: qs) =
+    match_patternT (p, q) andalso match_patternsT (ps, qs)
+fun match_ptype (PType (_, ps), PType (_, qs)) = match_patternsT (ps, qs)
 
 (* Is there a unifiable constant? *)
 fun pconst_mem f consts (s, ps) =
@@ -69,14 +132,9 @@
 fun pconst_hyper_mem f const_tab (s, ps) =
   exists (curry (match_ptype o f) ps) (these (Symtab.lookup const_tab s))
 
-fun pattern_of_type (Type (s, Ts)) = PApp (s, map pattern_of_type Ts)
-  | pattern_of_type (TFree (s, _)) = PApp (s, [])
-  | pattern_of_type (TVar _) = PVar
-
 (* Pairs a constant with the list of its type instantiations. *)
 fun ptype thy const x =
-  (if const then map pattern_of_type (these (try (Sign.const_typargs thy) x))
-   else [])
+  (if const then these (try (Sign.const_typargs thy) x) else [])
 fun rich_ptype thy const (s, T) =
   PType (order_of_type T, ptype thy const (s, T))
 fun rich_pconst thy const (s, T) = (s, rich_ptype thy const (s, T))
@@ -84,9 +142,19 @@
 fun string_of_hyper_pconst (s, ps) =
   s ^ "{" ^ commas (map string_of_ptype ps) ^ "}"
 
-(* Add a pconstant to the table, but a [] entry means a standard
-   connective, which we ignore.*)
-fun add_pconst_to_table (s, p) = Symtab.map_default (s, [p]) (insert (op =) p)
+fun patternT_eq (TVar _, TVar _) = true
+  | patternT_eq (Type (s, Ts), Type (t, Us)) = s = t andalso patternsT_eq (Ts, Us)
+  | patternT_eq (TFree (s, _), TFree (t, _)) = (s = t)
+  | patternT_eq _ = false
+and patternsT_eq ([], []) = true
+  | patternsT_eq ([], _) = false
+  | patternsT_eq (_, []) = false
+  | patternsT_eq (T :: Ts, U :: Us) = patternT_eq (T, U) andalso patternsT_eq (Ts, Us)
+fun ptype_eq (PType (m, Ts), PType (n, Us)) = m = n andalso patternsT_eq (Ts, Us)
+
+ (* Add a pconstant to the table, but a [] entry means a standard connective,
+    which we ignore. *)
+fun add_pconst_to_table (s, p) = Symtab.map_default (s, [p]) (insert ptype_eq p)
 
 (* Set constants tend to pull in too many irrelevant facts. We limit the damage
    by treating them more or less as if they were built-in but add their
@@ -94,21 +162,18 @@
 val set_consts = [@{const_name Collect}, @{const_name Set.member}]
 val set_thms = @{thms Collect_mem_eq mem_Collect_eq Collect_cong}
 
-fun add_pconsts_in_term thy is_built_in_const =
+fun add_pconsts_in_term thy =
   let
-    fun do_const const ext_arg (x as (s, _)) ts =
-      let val (built_in, ts) = is_built_in_const x ts in
-        if member (op =) set_consts s then
-          fold (do_term ext_arg) ts
-        else
-          (not built_in
-           ? add_pconst_to_table (rich_pconst thy const x))
-          #> fold (do_term false) ts
-      end
+    fun do_const const (x as (s, _)) ts =
+      if member (op =) set_consts s then
+        fold (do_term false) ts
+      else
+        (not (is_irrelevant_const s) ? add_pconst_to_table (rich_pconst thy const x))
+        #> fold (do_term false) ts
     and do_term ext_arg t =
       case strip_comb t of
-        (Const x, ts) => do_const true ext_arg x ts
-      | (Free x, ts) => do_const false ext_arg x ts
+        (Const x, ts) => do_const true x ts
+      | (Free x, ts) => do_const false x ts
       | (Abs (_, T, t'), ts) =>
         ((null ts andalso not ext_arg)
          (* Since lambdas on the right-hand side of equalities are usually
@@ -122,7 +187,7 @@
       if T = HOLogic.boolT then do_formula else do_term ext_arg
     and do_formula t =
       case t of
-        Const (@{const_name all}, _) $ Abs (_, T, t') => do_formula t'
+        Const (@{const_name all}, _) $ Abs (_, _, t') => do_formula t'
       | @{const "==>"} $ t1 $ t2 => do_formula t1 #> do_formula t2
       | Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 =>
         do_term_or_formula false T t1 #> do_term_or_formula true T t2
@@ -130,8 +195,8 @@
       | @{const False} => I
       | @{const True} => I
       | @{const Not} $ t1 => do_formula t1
-      | Const (@{const_name All}, _) $ Abs (_, T, t') => do_formula t'
-      | Const (@{const_name Ex}, _) $ Abs (_, T, t') => do_formula t'
+      | Const (@{const_name All}, _) $ Abs (_, _, t') => do_formula t'
+      | Const (@{const_name Ex}, _) $ Abs (_, _, t') => do_formula t'
       | @{const HOL.conj} $ t1 $ t2 => do_formula t1 #> do_formula t2
       | @{const HOL.disj} $ t1 $ t2 => do_formula t1 #> do_formula t2
       | @{const HOL.implies} $ t1 $ t2 => do_formula t1 #> do_formula t2
@@ -140,19 +205,19 @@
       | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])]))
         $ t1 $ t2 $ t3 =>
         do_formula t1 #> fold (do_term_or_formula false T) [t2, t3]
-      | Const (@{const_name Ex1}, _) $ Abs (_, T, t') => do_formula t'
-      | Const (@{const_name Ball}, _) $ t1 $ Abs (_, T, t') =>
+      | Const (@{const_name Ex1}, _) $ Abs (_, _, t') => do_formula t'
+      | Const (@{const_name Ball}, _) $ t1 $ Abs (_, _, t') =>
         do_formula (t1 $ Bound ~1) #> do_formula t'
-      | Const (@{const_name Bex}, _) $ t1 $ Abs (_, T, t') =>
+      | Const (@{const_name Bex}, _) $ t1 $ Abs (_, _, t') =>
         do_formula (t1 $ Bound ~1) #> do_formula t'
       | (t0 as Const (_, @{typ bool})) $ t1 =>
         do_term false t0 #> do_formula t1  (* theory constant *)
       | _ => do_term false t
   in do_formula end
 
-fun pconsts_in_fact thy is_built_in_const t =
+fun pconsts_in_fact thy t =
   Symtab.fold (fn (s, pss) => fold (cons o pair s) pss)
-              (Symtab.empty |> add_pconsts_in_term thy is_built_in_const t) []
+              (Symtab.empty |> add_pconsts_in_term thy t) []
 
 (* Inserts a dummy "constant" referring to the theory name, so that relevance
    takes the given theory into account. *)
@@ -167,9 +232,9 @@
 fun theory_const_prop_of fudge th =
   theory_constify fudge (Context.theory_name (theory_of_thm th)) (prop_of th)
 
-fun pair_consts_fact thy is_built_in_const fudge fact =
+fun pair_consts_fact thy fudge fact =
   case fact |> snd |> theory_const_prop_of fudge
-            |> pconsts_in_fact thy is_built_in_const of
+            |> pconsts_in_fact thy of
     [] => NONE
   | consts => SOME ((fact, consts), NONE)
 
@@ -177,15 +242,22 @@
    first by constant name and second by its list of type instantiations. For the
    latter, we need a linear ordering on "pattern list". *)
 
-fun pattern_ord p =
+fun patternT_ord p =
   case p of
-    (PVar, PVar) => EQUAL
-  | (PVar, PApp _) => LESS
-  | (PApp _, PVar) => GREATER
-  | (PApp q1, PApp q2) =>
-    prod_ord fast_string_ord (dict_ord pattern_ord) (q1, q2)
-fun ptype_ord (PType p, PType q) =
-  prod_ord (dict_ord pattern_ord) int_ord (swap p, swap q)
+    (Type (s, ps), Type (t, qs)) =>
+    (case fast_string_ord (s, t) of
+      EQUAL => dict_ord patternT_ord (ps, qs)
+    | ord => ord)
+  | (TVar _, TVar _) => EQUAL
+  | (TVar _, _) => LESS
+  | (Type _, TVar _) => GREATER
+  | (Type _, TFree _) => LESS
+  | (TFree (s, _), TFree (t, _)) => fast_string_ord (s, t)
+  | (TFree _, _) => GREATER
+fun ptype_ord (PType (m, ps), PType (n, qs)) =
+  (case dict_ord patternT_ord (ps, qs) of
+    EQUAL => int_ord (m, n)
+  | ord => ord)
 
 structure PType_Tab = Table(type key = ptype val ord = ptype_ord)
 
@@ -326,23 +398,23 @@
     (accepts, more_rejects @ rejects)
   end
 
-fun if_empty_replace_with_scope thy is_built_in_const facts sc tab =
+fun if_empty_replace_with_scope thy facts sc tab =
   if Symtab.is_empty tab then
     Symtab.empty
-    |> fold (add_pconsts_in_term thy is_built_in_const)
+    |> fold (add_pconsts_in_term thy)
             (map_filter (fn ((_, (sc', _)), th) =>
                             if sc' = sc then SOME (prop_of th) else NONE) facts)
   else
     tab
 
-fun consider_arities is_built_in_const th =
+fun consider_arities th =
   let
     fun aux _ _ NONE = NONE
       | aux t args (SOME tab) =
         case t of
           t1 $ t2 => SOME tab |> aux t1 (t2 :: args) |> aux t2 []
-        | Const (x as (s, _)) =>
-          (if is_built_in_const x args |> fst then
+        | Const (s, _) =>
+          (if is_widely_irrelevant_const s then
              SOME tab
            else case Symtab.lookup tab s of
              NONE => SOME (Symtab.update (s, length args) tab)
@@ -351,24 +423,24 @@
   in aux (prop_of th) [] end
 
 (* FIXME: This is currently only useful for polymorphic type encodings. *)
-fun could_benefit_from_ext is_built_in_const facts =
-  fold (consider_arities is_built_in_const o snd) facts (SOME Symtab.empty)
-  |> is_none
+fun could_benefit_from_ext facts =
+  fold (consider_arities o snd) facts (SOME Symtab.empty) |> is_none
 
 (* High enough so that it isn't wrongly considered as very relevant (e.g., for E
    weights), but low enough so that it is unlikely to be truncated away if few
    facts are included. *)
 val special_fact_index = 45 (* FUDGE *)
 
+fun eq_prod eqx eqy ((x1, y1), (x2, y2)) = eqx (x1, x2) andalso eqy (y1, y2)
+
 val really_hopeless_get_kicked_out_iter = 5 (* FUDGE *)
 
-fun relevance_filter ctxt thres0 decay max_facts is_built_in_const
-        (fudge as {threshold_divisor, ridiculous_threshold, ...}) facts hyp_ts
-        concl_t =
+fun relevance_filter ctxt thres0 decay max_facts
+        (fudge as {threshold_divisor, ridiculous_threshold, ...}) facts hyp_ts concl_t =
   let
     val thy = Proof_Context.theory_of ctxt
     val const_tab = fold (count_fact_consts thy fudge) facts Symtab.empty
-    val add_pconsts = add_pconsts_in_term thy is_built_in_const
+    val add_pconsts = add_pconsts_in_term thy
     val chained_ts =
       facts |> map_filter (fn ((_, (Chained, _)), th) => SOME (prop_of th)
                             | _ => NONE)
@@ -378,8 +450,7 @@
       |> fold add_pconsts hyp_ts
       |> add_pconsts concl_t
       |> (fn tab => if Symtab.is_empty tab then chained_const_tab else tab)
-      |> fold (if_empty_replace_with_scope thy is_built_in_const facts)
-              [Chained, Assum, Local]
+      |> fold (if_empty_replace_with_scope thy facts) [Chained, Assum, Local]
     fun iter j remaining_max thres rel_const_tab hopeless hopeful =
       let
         val hopeless =
@@ -421,7 +492,8 @@
               trace_msg ctxt (fn () => "New or updated constants: " ^
                   commas (rel_const_tab'
                           |> Symtab.dest
-                          |> subtract (op =) (rel_const_tab |> Symtab.dest)
+                          |> subtract (eq_prod (op =) (eq_list ptype_eq))
+                                      (rel_const_tab |> Symtab.dest)
                           |> map string_of_hyper_pconst));
               map (fst o fst) accepts @
               (if remaining_max = 0 then
@@ -475,41 +547,33 @@
         in bef @ add @ after end
     fun insert_special_facts accepts =
       (* FIXME: get rid of "ext" here once it is treated as a helper *)
-      [] |> could_benefit_from_ext is_built_in_const accepts ? cons @{thm ext}
+      [] |> could_benefit_from_ext accepts ? cons @{thm ext}
          |> add_set_const_thms accepts
          |> insert_into_facts accepts
   in
-    facts |> map_filter (pair_consts_fact thy is_built_in_const fudge)
+    facts |> map_filter (pair_consts_fact thy fudge)
           |> iter 0 max_facts thres0 goal_const_tab []
           |> insert_special_facts
           |> tap (fn accepts => trace_msg ctxt (fn () =>
                       "Total relevant: " ^ string_of_int (length accepts)))
   end
 
-fun mepo_suggested_facts ctxt
-        ({fact_thresholds = (thres0, thres1), ...} : params) prover
-        max_facts fudge hyp_ts concl_t facts =
+fun mepo_suggested_facts ctxt ({fact_thresholds = (thres0, thres1), ...} : params) max_facts fudge
+      hyp_ts concl_t facts =
   let
     val thy = Proof_Context.theory_of ctxt
-    val is_built_in_const =
-      Sledgehammer_Provers.is_built_in_const_of_prover ctxt prover
-    val fudge =
-      case fudge of
-        SOME fudge => fudge
-      | NONE => Sledgehammer_Provers.relevance_fudge_of_prover ctxt prover
+    val fudge = fudge |> the_default default_relevance_fudge
     val decay = Math.pow ((1.0 - thres1) / (1.0 - thres0),
                           1.0 / Real.fromInt (max_facts + 1))
   in
-    trace_msg ctxt (fn () => "Considering " ^ string_of_int (length facts) ^
-                             " facts");
+    trace_msg ctxt (fn () => "Considering " ^ string_of_int (length facts) ^ " facts");
     (if thres1 < 0.0 then
        facts
-     else if thres0 > 1.0 orelse thres0 > thres1 then
+     else if thres0 > 1.0 orelse thres0 > thres1 orelse max_facts <= 0 then
        []
      else
-       relevance_filter ctxt thres0 decay max_facts is_built_in_const fudge
-           facts hyp_ts
-           (concl_t |> theory_constify fudge (Context.theory_name thy)))
+       relevance_filter ctxt thres0 decay max_facts fudge facts hyp_ts
+         (concl_t |> theory_constify fudge (Context.theory_name thy)))
     |> map fact_of_raw_fact
   end
 
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -18,7 +18,7 @@
   val auto_minimize_max_time : real Config.T
   val minimize_facts :
     (string -> thm list -> unit) -> string -> params -> bool -> int -> int
-    -> Proof.state -> play Lazy.lazy option
+    -> Proof.state -> thm -> play Lazy.lazy option
     -> ((string * stature) * thm list) list
     -> ((string * stature) * thm list) list option
        * (play Lazy.lazy * (play -> string) * string)
@@ -59,7 +59,7 @@
                  max_new_mono_instances, type_enc, strict, lam_trans,
                  uncurried_aliases, isar_proofs, isar_compress, isar_try0,
                  preplay_timeout, ...} : params)
-               silent (prover : prover) timeout i n state facts =
+               silent (prover : prover) timeout i n state goal facts =
   let
     val _ =
       print silent (fn () =>
@@ -70,7 +70,6 @@
              | _ => ""
            else
              "") ^ "...")
-    val {goal, ...} = Proof.goal state
     val facts = facts |> maps (fn (n, ths) => map (pair n) ths)
     val params =
       {debug = debug, verbose = verbose, overlord = overlord, spy = spy, blocking = true,
@@ -82,7 +81,8 @@
        slice = false, minimize = SOME false, timeout = timeout, preplay_timeout = preplay_timeout,
        expect = ""}
     val problem =
-      {state = state, goal = goal, subgoal = i, subgoal_count = n, factss = [("", facts)]}
+      {comment = "", state = state, goal = goal, subgoal = i, subgoal_count = n,
+       factss = [("", facts)]}
     val result as {outcome, used_facts, run_time, ...} =
       prover params (K (K (K ""))) problem
   in
@@ -191,12 +191,12 @@
   end
 
 fun minimize_facts do_learn prover_name (params as {learn, timeout, ...}) silent
-                   i n state preplay0 facts =
+                   i n state goal preplay0 facts =
   let
     val ctxt = Proof.context_of state
     val prover =
       get_prover ctxt (if silent then Auto_Minimize else Minimize) prover_name
-    fun test timeout = test_facts params silent prover timeout i n state
+    fun test timeout = test_facts params silent prover timeout i n state goal
     val (chained, non_chained) = List.partition is_fact_chained facts
     (* Push chained facts to the back, so that they are less likely to be
        kicked out by the linear minimization algorithm. *)
@@ -272,7 +272,7 @@
 
 fun maybe_minimize ctxt mode do_learn name
         (params as {verbose, isar_proofs, minimize, ...})
-        ({state, subgoal, subgoal_count, ...} : prover_problem)
+        ({state, goal, subgoal, subgoal_count, ...} : prover_problem)
         (result as {outcome, used_facts, used_from, run_time, preplay, message,
                     message_tail} : prover_result) =
   if is_some outcome orelse null used_facts then
@@ -319,8 +319,8 @@
       val (used_facts, (preplay, message, _)) =
         if minimize then
           minimize_facts do_learn minimize_name params
-              (mode <> Normal orelse not verbose) subgoal subgoal_count state
-              (SOME preplay)
+              (not verbose orelse (mode <> Normal andalso mode <> MaSh)) subgoal
+              subgoal_count state goal (SOME preplay)
               (filter_used_facts true used_facts (map (apsnd single) used_from))
           |>> Option.map (map fst)
         else
@@ -342,12 +342,10 @@
 fun run_minimize (params as {provers, ...}) do_learn i refs state =
   let
     val ctxt = Proof.context_of state
+    val {goal, facts = chained_ths, ...} = Proof.goal state
     val reserved = reserved_isar_keyword_table ()
-    val chained_ths = #facts (Proof.goal state)
     val css = clasimpset_rule_table_of ctxt
-    val facts =
-      refs |> maps (map (apsnd single)
-                    o fact_of_ref ctxt reserved chained_ths css)
+    val facts = refs |> maps (map (apsnd single) o fact_of_ref ctxt reserved chained_ths css)
   in
     case subgoal_count state of
       0 => Output.urgent_message "No subgoal!"
@@ -355,7 +353,7 @@
              [] => error "No prover is set."
            | prover :: _ =>
              (kill_provers ();
-              minimize_facts do_learn prover params false i n state NONE facts
+              minimize_facts do_learn prover params false i n state goal NONE facts
               |> (fn (_, (preplay, message, message_tail)) =>
                      message (Lazy.force preplay) ^ message_tail
                      |> Output.urgent_message))
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -44,28 +44,9 @@
      preplay_timeout : Time.time option,
      expect : string}
 
-  type relevance_fudge =
-    {local_const_multiplier : real,
-     worse_irrel_freq : real,
-     higher_order_irrel_weight : real,
-     abs_rel_weight : real,
-     abs_irrel_weight : real,
-     theory_const_rel_weight : real,
-     theory_const_irrel_weight : real,
-     chained_const_irrel_weight : real,
-     intro_bonus : real,
-     elim_bonus : real,
-     simp_bonus : real,
-     local_bonus : real,
-     assum_bonus : real,
-     chained_bonus : real,
-     max_imperfect : real,
-     max_imperfect_exp : real,
-     threshold_divisor : real,
-     ridiculous_threshold : real}
-
   type prover_problem =
-    {state : Proof.state,
+    {comment : string,
+     state : Proof.state,
      goal : thm,
      subgoal : int,
      subgoal_count : int,
@@ -88,8 +69,7 @@
   val problem_prefix : string Config.T
   val completish : bool Config.T
   val atp_full_names : bool Config.T
-  val smt_builtin_facts : bool Config.T
-  val smt_builtin_trans : bool Config.T
+  val smt_builtins : bool Config.T
   val smt_triggers : bool Config.T
   val smt_weights : bool Config.T
   val smt_weight_min_facts : int Config.T
@@ -117,14 +97,9 @@
     Proof.context -> string -> string option
   val remotify_prover_if_not_installed :
     Proof.context -> string -> string option
-  val default_max_facts_of_prover : Proof.context -> bool -> string -> int
+  val default_max_facts_of_prover : Proof.context -> string -> int
   val is_unit_equality : term -> bool
   val is_appropriate_prop_of_prover : Proof.context -> string -> term -> bool
-  val is_built_in_const_of_prover :
-    Proof.context -> string -> string * typ -> term list -> bool * term list
-  val atp_relevance_fudge : relevance_fudge
-  val smt_relevance_fudge : relevance_fudge
-  val relevance_fudge_of_prover : Proof.context -> string -> relevance_fudge
   val weight_smt_fact :
     Proof.context -> int -> ((string * stature) * thm) * int
     -> (string * stature) * (int option * thm)
@@ -160,27 +135,18 @@
 (** The Sledgehammer **)
 
 (* Empty string means create files in Isabelle's temporary files directory. *)
-val dest_dir =
-  Attrib.setup_config_string @{binding sledgehammer_dest_dir} (K "")
-val problem_prefix =
-  Attrib.setup_config_string @{binding sledgehammer_problem_prefix} (K "prob")
-val completish =
-  Attrib.setup_config_bool @{binding sledgehammer_completish} (K false)
+val dest_dir = Attrib.setup_config_string @{binding sledgehammer_dest_dir} (K "")
+val problem_prefix = Attrib.setup_config_string @{binding sledgehammer_problem_prefix} (K "prob")
+val completish = Attrib.setup_config_bool @{binding sledgehammer_completish} (K false)
 
 (* In addition to being easier to read, readable names are often much shorter,
    especially if types are mangled in names. This makes a difference for some
    provers (e.g., E). For these reason, short names are enabled by default. *)
-val atp_full_names =
-  Attrib.setup_config_bool @{binding sledgehammer_atp_full_names} (K false)
+val atp_full_names = Attrib.setup_config_bool @{binding sledgehammer_atp_full_names} (K false)
 
-val smt_builtin_facts =
-  Attrib.setup_config_bool @{binding sledgehammer_smt_builtin_facts} (K true)
-val smt_builtin_trans =
-  Attrib.setup_config_bool @{binding sledgehammer_smt_builtin_trans} (K true)
-val smt_triggers =
-  Attrib.setup_config_bool @{binding sledgehammer_smt_triggers} (K true)
-val smt_weights =
-  Attrib.setup_config_bool @{binding sledgehammer_smt_weights} (K true)
+val smt_builtins = Attrib.setup_config_bool @{binding sledgehammer_smt_builtins} (K true)
+val smt_triggers = Attrib.setup_config_bool @{binding sledgehammer_smt_triggers} (K true)
+val smt_weights = Attrib.setup_config_bool @{binding sledgehammer_smt_weights} (K true)
 val smt_weight_min_facts =
   Attrib.setup_config_int @{binding sledgehammer_smt_weight_min_facts} (K 20)
 
@@ -240,15 +206,14 @@
 
 val reconstructor_default_max_facts = 20
 
-fun slice_max_facts (_, (_, ( ((max_facts, _), _, _, _, _), _))) = max_facts
+fun slice_max_facts (_, ( ((max_facts, _), _, _, _, _), _)) = max_facts
 
-fun default_max_facts_of_prover ctxt slice name =
+fun default_max_facts_of_prover ctxt name =
   let val thy = Proof_Context.theory_of ctxt in
     if is_reconstructor name then
       reconstructor_default_max_facts
     else if is_atp thy name then
-      fold (Integer.max o slice_max_facts)
-           (get_slices slice (#best_slices (get_atp thy name ()) ctxt)) 0
+      fold (Integer.max o slice_max_facts) (#best_slices (get_atp thy name ()) ctxt) 0
     else (* is_smt_prover ctxt name *)
       SMT_Solver.default_max_relevant ctxt name
   end
@@ -275,68 +240,6 @@
 fun is_appropriate_prop_of_prover ctxt name =
   if is_unit_equational_atp ctxt name then is_unit_equality else K true
 
-val atp_irrelevant_const_tab =
-  Symtab.make (map (rpair ()) atp_irrelevant_consts)
-
-fun is_built_in_const_of_prover ctxt name =
-  if is_smt_prover ctxt name andalso Config.get ctxt smt_builtin_facts then
-    let val ctxt = ctxt |> select_smt_solver name in
-      fn x => fn ts =>
-         if SMT_Builtin.is_builtin_num_ext ctxt (list_comb (Const x, ts)) then
-           (true, [])
-         else if SMT_Builtin.is_builtin_fun_ext ctxt x ts then
-           (true, ts)
-         else
-           (false, ts)
-    end
-  else
-    fn (s, _) => fn ts => (Symtab.defined atp_irrelevant_const_tab s, ts)
-
-(* FUDGE *)
-val atp_relevance_fudge =
-  {local_const_multiplier = 1.5,
-   worse_irrel_freq = 100.0,
-   higher_order_irrel_weight = 1.05,
-   abs_rel_weight = 0.5,
-   abs_irrel_weight = 2.0,
-   theory_const_rel_weight = 0.5,
-   theory_const_irrel_weight = 0.25,
-   chained_const_irrel_weight = 0.25,
-   intro_bonus = 0.15,
-   elim_bonus = 0.15,
-   simp_bonus = 0.15,
-   local_bonus = 0.55,
-   assum_bonus = 1.05,
-   chained_bonus = 1.5,
-   max_imperfect = 11.5,
-   max_imperfect_exp = 1.0,
-   threshold_divisor = 2.0,
-   ridiculous_threshold = 0.01}
-
-(* FUDGE (FIXME) *)
-val smt_relevance_fudge =
-  {local_const_multiplier = #local_const_multiplier atp_relevance_fudge,
-   worse_irrel_freq = #worse_irrel_freq atp_relevance_fudge,
-   higher_order_irrel_weight = #higher_order_irrel_weight atp_relevance_fudge,
-   abs_rel_weight = #abs_rel_weight atp_relevance_fudge,
-   abs_irrel_weight = #abs_irrel_weight atp_relevance_fudge,
-   theory_const_rel_weight = #theory_const_rel_weight atp_relevance_fudge,
-   theory_const_irrel_weight = #theory_const_irrel_weight atp_relevance_fudge,
-   chained_const_irrel_weight = #chained_const_irrel_weight atp_relevance_fudge,
-   intro_bonus = #intro_bonus atp_relevance_fudge,
-   elim_bonus = #elim_bonus atp_relevance_fudge,
-   simp_bonus = #simp_bonus atp_relevance_fudge,
-   local_bonus = #local_bonus atp_relevance_fudge,
-   assum_bonus = #assum_bonus atp_relevance_fudge,
-   chained_bonus = #chained_bonus atp_relevance_fudge,
-   max_imperfect = #max_imperfect atp_relevance_fudge,
-   max_imperfect_exp = #max_imperfect_exp atp_relevance_fudge,
-   threshold_divisor = #threshold_divisor atp_relevance_fudge,
-   ridiculous_threshold = #ridiculous_threshold atp_relevance_fudge}
-
-fun relevance_fudge_of_prover ctxt name =
-  if is_smt_prover ctxt name then smt_relevance_fudge else atp_relevance_fudge
-
 fun supported_provers ctxt =
   let
     val thy = Proof_Context.theory_of ctxt
@@ -383,28 +286,9 @@
    preplay_timeout : Time.time option,
    expect : string}
 
-type relevance_fudge =
-  {local_const_multiplier : real,
-   worse_irrel_freq : real,
-   higher_order_irrel_weight : real,
-   abs_rel_weight : real,
-   abs_irrel_weight : real,
-   theory_const_rel_weight : real,
-   theory_const_irrel_weight : real,
-   chained_const_irrel_weight : real,
-   intro_bonus : real,
-   elim_bonus : real,
-   simp_bonus : real,
-   local_bonus : real,
-   assum_bonus : real,
-   chained_bonus : real,
-   max_imperfect : real,
-   max_imperfect_exp : real,
-   threshold_divisor : real,
-   ridiculous_threshold : real}
-
 type prover_problem =
-  {state : Proof.state,
+  {comment : string,
+   state : Proof.state,
    goal : thm,
    subgoal : int,
    subgoal_count : int,
@@ -681,7 +565,7 @@
                     max_new_mono_instances, isar_proofs, isar_compress,
                     isar_try0, slice, timeout, preplay_timeout, ...})
         minimize_command
-        ({state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
+        ({comment, state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
   let
     val thy = Proof.theory_of state
     val ctxt = Proof.context_of state
@@ -734,15 +618,15 @@
       let
         (* If slicing is disabled, we expand the last slice to fill the entire
            time available. *)
-        val actual_slices = get_slices slice (best_slices ctxt)
+        val all_slices = best_slices ctxt
+        val actual_slices = get_slices slice all_slices
+        fun max_facts_of_slices f slices = fold (Integer.max o slice_max_facts o f) slices 0
         val num_actual_slices = length actual_slices
         val max_fact_factor =
-          case max_facts of
-            NONE => 1.0
-          | SOME max =>
-            Real.fromInt max
-            / Real.fromInt (fold (Integer.max o slice_max_facts)
-                                 actual_slices 0)
+          Real.fromInt (case max_facts of
+              NONE => max_facts_of_slices I all_slices
+            | SOME max => max)
+          / Real.fromInt (max_facts_of_slices snd actual_slices)
         fun monomorphize_facts facts =
           let
             val ctxt =
@@ -845,7 +729,8 @@
             val _ =
               atp_problem
               |> lines_of_atp_problem format ord ord_info
-              |> cons ("% " ^ command ^ "\n")
+              |> cons ("% " ^ command ^ "\n" ^
+                (if comment = "" then "" else "% " ^ comment ^ "\n"))
               |> File.write_list prob_path
             val ((output, run_time), (atp_proof, outcome)) =
               time_limit generous_slice_timeout Isabelle_System.bash_output
@@ -1028,10 +913,8 @@
 val is_boring_builtin_typ =
   not o exists_subtype (member (op =) [@{typ nat}, @{typ int}, HOLogic.realT])
 
-fun smt_filter_loop name
-                    ({debug, verbose, overlord, max_mono_iters,
-                      max_new_mono_instances, timeout, slice, ...} : params)
-                    state goal i =
+fun smt_filter_loop name ({debug, overlord, max_mono_iters, max_new_mono_instances, timeout, slice,
+      ...} : params) state goal i =
   let
     fun repair_context ctxt =
       ctxt |> select_smt_solver name
@@ -1044,7 +927,7 @@
                  I)
            |> Config.put SMT_Config.infer_triggers
                          (Config.get ctxt smt_triggers)
-           |> not (Config.get ctxt smt_builtin_trans)
+           |> not (Config.get ctxt smt_builtins)
               ? (SMT_Builtin.filter_builtins is_boring_builtin_typ
                  #> Config.put SMT_Config.datatypes false)
            |> repair_monomorph_context max_mono_iters default_max_mono_iters
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -640,7 +640,7 @@
 fun isar_proof_would_be_a_good_idea preplay =
   case preplay of
     Played (reconstr, _) => reconstr = SMT
-  | Trust_Playable _ => true
+  | Trust_Playable _ => false
   | Failed_to_Play _ => true
 
 fun proof_text ctxt isar_proofs isar_params num_chained
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_run.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_run.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -68,21 +68,21 @@
 fun launch_prover (params as {debug, verbose, spy, blocking, max_facts, slice,
                               timeout, expect, ...})
         mode output_result minimize_command only learn
-        {state, goal, subgoal, subgoal_count, factss as (_, facts) :: _} name =
+        {comment, state, goal, subgoal, subgoal_count, factss as (_, facts) :: _} name =
   let
     val ctxt = Proof.context_of state
 
     val hard_timeout = time_mult 3.0 (timeout |> the_default one_day)
-    val _ = spying spy (fn () => (state, subgoal, name, "launched"));
+    val _ = spying spy (fn () => (state, subgoal, name, "Launched"));
     val birth_time = Time.now ()
     val death_time = Time.+ (birth_time, hard_timeout)
-    val max_facts = max_facts |> the_default (default_max_facts_of_prover ctxt slice name)
+    val max_facts = max_facts |> the_default (default_max_facts_of_prover ctxt name)
     val num_facts = length facts |> not only ? Integer.min max_facts
 
     fun desc () = prover_description ctxt params name num_facts subgoal subgoal_count goal
 
     val problem =
-      {state = state, goal = goal, subgoal = subgoal,
+      {comment = comment, state = state, goal = goal, subgoal = subgoal,
        subgoal_count = subgoal_count,
        factss = factss
          |> map (apsnd ((not (is_ho_atp ctxt name)
@@ -100,13 +100,33 @@
                   " proof (of " ^ string_of_int (length facts) ^ "): ") "."
       |> Output.urgent_message
 
-    fun spying_str_of_res ({outcome = NONE, used_facts, ...} : prover_result) =
-        let val num_used_facts = length used_facts in
-          "success: Found proof with " ^ string_of_int num_used_facts ^ " fact" ^
-          plural_s num_used_facts
+    fun spying_str_of_res ({outcome = NONE, used_facts, used_from, ...} : prover_result) =
+        let
+          val num_used_facts = length used_facts
+
+          fun find_indices facts =
+            tag_list 1 facts
+            |> map (fn (j, fact) => fact |> apsnd (K j))
+            |> filter_used_facts false used_facts
+            |> map (prefix "@" o string_of_int o snd)
+
+          fun filter_info (fact_filter, facts) =
+            let
+              val indices = find_indices facts
+              val unknowns = replicate (num_used_facts - length indices) "?"
+            in (commas (indices @ unknowns), fact_filter) end
+
+          val filter_infos =
+            map filter_info (("actual", used_from) :: factss)
+            |> AList.group (op =)
+            |> map (fn (indices, fact_filters) => commas fact_filters ^ ": " ^ indices)
+        in
+          "Success: Found proof with " ^ string_of_int num_used_facts ^
+          " of " ^ string_of_int num_facts ^ " fact" ^ plural_s num_facts ^
+          (if num_used_facts = 0 then "" else ": " ^ commas filter_infos)
         end
       | spying_str_of_res {outcome = SOME failure, ...} =
-        "failure: " ^ string_of_atp_failure failure
+        "Failure: " ^ string_of_atp_failure failure
 
     fun really_go () =
       problem
@@ -209,8 +229,7 @@
       val _ = Proof.assert_backward state
       val print =
         if mode = Normal andalso is_none output_result then Output.urgent_message else K ()
-      val state =
-        state |> Proof.map_context (Config.put SMT_Config.verbose debug)
+      val state = state |> Proof.map_context (Config.put SMT_Config.verbose debug)
       val ctxt = Proof.context_of state
       val {facts = chained, goal, ...} = Proof.goal state
       val (_, hyp_ts, concl_t) = strip_subgoal goal i ctxt
@@ -225,11 +244,9 @@
                 SOME name => error ("No such prover: " ^ name ^ ".")
               | NONE => ()
       val _ = print "Sledgehammering..."
-      val _ =
-        spying spy (fn () => (state, i, "***", "starting " ^ @{make_string} mode ^ " mode"))
-      val (smts, (ueq_atps, full_atps)) =
-        provers |> List.partition (is_smt_prover ctxt)
-                ||> List.partition (is_unit_equational_atp ctxt)
+      val _ = spying spy (fn () => (state, i, "***", "Starting " ^ @{make_string} mode ^ " mode"))
+
+      val (ueq_atps, full_provers) = List.partition (is_unit_equational_atp ctxt) provers
 
       val spying_str_of_factss =
         commas o map (fn (filter, facts) => filter ^ ": " ^ string_of_int (length facts))
@@ -240,18 +257,16 @@
             case max_facts of
               SOME n => n
             | NONE =>
-              0 |> fold (Integer.max o default_max_facts_of_prover ctxt slice)
-                        provers
+              0 |> fold (Integer.max o default_max_facts_of_prover ctxt) provers
                 |> mode = Auto_Try ? (fn n => n div auto_try_max_facts_divisor)
           val _ = spying spy (fn () => (state, i, label ^ "s",
-            "filtering " ^ string_of_int (length all_facts) ^ " facts"));
+            "Filtering " ^ string_of_int (length all_facts) ^ " facts"));
         in
           all_facts
           |> (case is_appropriate_prop of
                 SOME is_app => filter (is_app o prop_of o snd)
               | NONE => I)
-          |> relevant_facts ctxt params (hd provers) max_max_facts fact_override
-                            hyp_ts concl_t
+          |> relevant_facts ctxt params (hd provers) max_max_facts fact_override hyp_ts concl_t
           |> tap (fn factss =>
                      if verbose then
                        label ^ plural_s (length provers) ^ ": " ^
@@ -259,16 +274,15 @@
                        |> print
                      else
                        ())
-          |> spy ? tap (fn factss =>
-            spying spy (fn () =>
-              (state, i, label ^ "s", "selected facts: " ^ spying_str_of_factss factss)))
+          |> spy ? tap (fn factss => spying spy (fn () =>
+            (state, i, label ^ "s", "Selected facts: " ^ spying_str_of_factss factss)))
         end
 
       fun launch_provers state label is_appropriate_prop provers =
         let
           val factss = get_factss label is_appropriate_prop provers
           val problem =
-            {state = state, goal = goal, subgoal = i, subgoal_count = n,
+            {comment = "", state = state, goal = goal, subgoal = i, subgoal_count = n,
              factss = factss}
           fun learn prover =
             mash_learn_proof ctxt params prover (prop_of goal) all_facts
@@ -287,30 +301,26 @@
             |> max_outcome_code |> rpair state
         end
 
-      fun launch_atps label is_appropriate_prop atps accum =
-        if null atps then
+      fun maybe_launch_provers label is_appropriate_prop provers_to_launch accum =
+        if null provers_to_launch then
           accum
         else if is_some is_appropriate_prop andalso
                 not (the is_appropriate_prop concl_t) then
-          (if verbose orelse length atps = length provers then
+          (if verbose orelse length provers_to_launch = length provers then
              "Goal outside the scope of " ^
-             space_implode " " (serial_commas "and" (map quote atps)) ^ "."
+             space_implode " " (serial_commas "and" (map quote provers_to_launch)) ^ "."
              |> Output.urgent_message
            else
              ();
            accum)
         else
-          launch_provers state label is_appropriate_prop atps
+          launch_provers state label is_appropriate_prop provers_to_launch
 
-      fun launch_smts accum =
-        if null smts then accum else launch_provers state "SMT solver" NONE smts
-
-      val launch_full_atps = launch_atps "ATP" NONE full_atps
-
-      val launch_ueq_atps = launch_atps "Unit-equational provers" (SOME is_unit_equality) ueq_atps
+      val launch_full_provers = maybe_launch_provers "ATP/SMT" NONE full_provers
+      val launch_ueq_atps = maybe_launch_provers "Unit-equational provers" (SOME is_unit_equality) ueq_atps
 
       fun launch_atps_and_smt_solvers p =
-        [launch_full_atps, launch_smts, launch_ueq_atps]
+        [launch_full_provers, launch_ueq_atps]
         |> Par_List.map (fn f => fst (f p))
         handle ERROR msg => (print ("Error: " ^ msg); error msg)
 
@@ -318,14 +328,9 @@
         accum |> (mode = Normal orelse outcome_code <> someN) ? f
     in
       (unknownN, state)
-      |> (if mode = Auto_Try then
-            launch_full_atps
-          else if blocking then
-            launch_atps_and_smt_solvers #> max_outcome_code #> rpair state
-          else
-            (fn p => (Future.fork (tap (fn () => launch_atps_and_smt_solvers p)); p)))
-      handle TimeLimit.TimeOut =>
-             (print "Sledgehammer ran out of time."; (unknownN, state))
+      |> (if blocking then launch_full_provers
+          else (fn p => (Future.fork (tap (fn () => launch_full_provers p)); p)))
+      handle TimeLimit.TimeOut => (print "Sledgehammer ran out of time."; (unknownN, state))
     end
     |> `(fn (outcome_code, _) => outcome_code = someN)
 
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -97,44 +97,29 @@
 (* FIXME: Similar yet different code in "mirabelle.ML". The code here has a few
    fixes that seem to be missing over there; or maybe the two code portions are
    not doing the same? *)
-fun fold_body_thms outer_name (map_plain_name, map_inclass_name) =
+fun fold_body_thm outer_name (map_plain_name, map_inclass_name) =
   let
-    fun app map_name n (PBody {thms, ...}) =
-      thms |> fold (fn (_, (name, _, body)) => fn accum =>
-          let
-            val collect = union (op =) o the_list o map_name
-            (* The "name = outer_name" case caters for the uncommon case where
-               the proved theorem occurs in its own proof (e.g.,
-               "Transitive_Closure.trancl_into_trancl"). *)
-            val (anonymous, enter_class) =
-              if name = "" orelse (n = 1 andalso name = outer_name) then
-                (true, false)
-              else if n = 1 andalso map_inclass_name name = SOME outer_name then
-                (true, true)
-              else
-                (false, false)
-            val accum =
-              accum |> (if n = 1 andalso not anonymous then collect name else I)
-            val n = n + (if anonymous then 0 else 1)
-          in
-            accum
-            |> (if n <= 1 then
-                  app (if enter_class then map_inclass_name else map_name) n
-                      (Future.join body)
-                else
-                  I)
-          end)
-  in fold (app map_plain_name 0) end
+    fun app_thm map_name (_, (name, _, body)) accum =
+      let
+        val (anonymous, enter_class) =
+          (* The "name = outer_name" case caters for the uncommon case where the proved theorem
+             occurs in its own proof (e.g., "Transitive_Closure.trancl_into_trancl"). *)
+          if name = "" orelse name = outer_name then (true, false)
+          else if map_inclass_name name = SOME outer_name then (true, true)
+          else (false, false)
+      in
+        if anonymous then
+          accum |> app_body (if enter_class then map_inclass_name else map_name) (Future.join body)
+        else
+          accum |> union (op =) (the_list (map_name name))
+      end
+    and app_body map_name (PBody {thms, ...}) = fold (app_thm map_name) thms
+  in app_body map_plain_name end
 
 fun thms_in_proof name_tabs th =
-  let
-    val map_names =
-      case name_tabs of
-        SOME p => pairself Symtab.lookup p
-      | NONE => `I SOME
-    val names =
-      fold_body_thms (Thm.get_name_hint th) map_names [Thm.proof_body_of th] []
-  in names end
+  let val map_names = (case name_tabs of SOME p => pairself Symtab.lookup p | NONE => `I SOME) in
+    fold_body_thm (Thm.get_name_hint th) map_names (Proofterm.strip_thm (Thm.proof_body_of th)) []
+  end
 
 fun thms_of_name ctxt name =
   let
@@ -162,7 +147,7 @@
 fun hackish_string_of_term ctxt =
   with_vanilla_print_mode (Syntax.string_of_term ctxt) #> simplify_spaces
 
-val spying_version = "b"
+val spying_version = "c"
 
 fun spying false _ = ()
   | spying true f =
--- a/src/HOL/Tools/groebner.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/groebner.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -21,11 +21,6 @@
 structure Groebner : GROEBNER =
 struct
 
-fun is_comb ct =
-  (case Thm.term_of ct of
-    _ $ _ => true
-  | _ => false);
-
 val concl = Thm.cprop_of #> Thm.dest_arg;
 
 fun is_binop ct ct' =
@@ -37,8 +32,6 @@
   if is_binop ct ct' then Thm.dest_binop ct'
   else raise CTERM ("dest_binary: bad binop", [ct, ct'])
 
-fun inst_thm inst = Thm.instantiate ([], inst);
-
 val rat_0 = Rat.zero;
 val rat_1 = Rat.one;
 val minus_rat = Rat.neg;
@@ -77,10 +70,6 @@
     n1 < n2 orelse n1 = n2 andalso lexorder m1 m2
     end;
 
-fun morder_le m1 m2 = morder_lt m1 m2 orelse (m1 = m2);
-
-fun morder_gt m1 m2 = morder_lt m2 m1;
-
 (* Arithmetic on canonical polynomials. *)
 
 fun grob_neg l = map (fn (c,m) => (minus_rat c,m)) l;
@@ -125,33 +114,9 @@
 
 fun grob_pow vars l n =
   if n < 0 then error "grob_pow: negative power"
-  else if n = 0 then [(rat_1,map (fn v => 0) vars)]
+  else if n = 0 then [(rat_1,map (K 0) vars)]
   else grob_mul l (grob_pow vars l (n - 1));
 
-fun degree vn p =
- case p of
-  [] => error "Zero polynomial"
-| [(c,ns)] => nth ns vn
-| (c,ns)::p' => Int.max (nth ns vn, degree vn p');
-
-fun head_deg vn p = let val d = degree vn p in
- (d,fold (fn (c,r) => fn q => grob_add q [(c, map_index (fn (i,n) => if i = vn then 0 else n) r)]) (filter (fn (c,ns) => c <>/ rat_0 andalso nth ns vn = d) p) []) end;
-
-val is_zerop = forall (fn (c,ns) => c =/ rat_0 andalso forall (curry (op =) 0) ns);
-val grob_pdiv =
- let fun pdiv_aux vn (n,a) p k s =
-  if is_zerop s then (k,s) else
-  let val (m,b) = head_deg vn s
-  in if m < n then (k,s) else
-     let val p' = grob_mul p [(rat_1, map_index (fn (i,v) => if i = vn then m - n else 0)
-                                                (snd (hd s)))]
-     in if a = b then pdiv_aux vn (n,a) p k (grob_sub s p')
-        else pdiv_aux vn (n,a) p (k + 1) (grob_sub (grob_mul a s) (grob_mul b p'))
-     end
-  end
- in fn vn => fn s => fn p => pdiv_aux vn (head_deg vn p) p 0 s
- end;
-
 (* Monomial division operation. *)
 
 fun mdiv (c1,m1) (c2,m2) =
@@ -160,7 +125,7 @@
 
 (* Lowest common multiple of two monomials. *)
 
-fun mlcm (c1,m1) (c2,m2) = (rat_1, ListPair.map Int.max (m1, m2));
+fun mlcm (_,m1) (_,m2) = (rat_1, ListPair.map Int.max (m1, m2));
 
 (* Reduce monomial cm by polynomial pol, returning replacement for cm.  *)
 
@@ -200,8 +165,8 @@
 
 fun spoly cm ph1 ph2 =
   case (ph1,ph2) of
-    (([],h),p) => ([],h)
-  | (p,([],h)) => ([],h)
+    (([],h),_) => ([],h)
+  | (_,([],h)) => ([],h)
   | ((cm1::ptl1,his1),(cm2::ptl2,his2)) =>
         (grob_sub (grob_cmul (mdiv cm cm1) ptl1)
                   (grob_cmul (mdiv cm cm2) ptl2),
@@ -218,12 +183,12 @@
 
 (* The most popular heuristic is to order critical pairs by LCM monomial.    *)
 
-fun forder ((c1,m1),_) ((c2,m2),_) = morder_lt m1 m2;
+fun forder ((_,m1),_) ((_,m2),_) = morder_lt m1 m2;
 
 fun poly_lt  p q =
   case (p,q) of
-    (p,[]) => false
-  | ([],q) => true
+    (_,[]) => false
+  | ([],_) => true
   | ((c1,m1)::o1,(c2,m2)::o2) =>
         c1 </ c2 orelse
         c1 =/ c2 andalso ((morder_lt m1 m2) orelse m1 = m2 andalso poly_lt o1 o2);
@@ -234,7 +199,7 @@
 fun poly_eq p1 p2 =
   eq_list (fn ((c1, m1), (c2, m2)) => c1 =/ c2 andalso (m1: int list) = m2) (p1, p2);
 
-fun memx ((p1,h1),(p2,h2)) ppairs =
+fun memx ((p1,_),(p2,_)) ppairs =
   not (exists (fn ((q1,_),(q2,_)) => poly_eq p1 q1 andalso poly_eq p2 q2) ppairs);
 
 (* Buchberger's second criterion.                                            *)
@@ -277,7 +242,7 @@
  case pairs of
    [] => basis
  | (l,(p1,p2))::opairs =>
-   let val (sph as (sp,hist)) = monic (reduce basis (spoly l p1 p2))
+   let val (sph as (sp,_)) = monic (reduce basis (spoly l p1 p2))
    in
     if null sp orelse criterion2 basis (l,(p1,p2)) opairs
     then grobner_basis basis opairs
@@ -324,7 +289,7 @@
 
 fun grobner_refute pols =
   let val gb = grobner pols in
-  snd(find (fn (p,h) => length p = 1 andalso forall (fn x=> x=0) (snd(hd p))) gb)
+  snd(find (fn (p,_) => length p = 1 andalso forall (fn x=> x=0) (snd(hd p))) gb)
   end;
 
 (* Turn proof into a certificate as sum of multipliers.                      *)
@@ -366,8 +331,8 @@
 
 fun grobner_strong vars pols pol =
     let val vars' = @{cterm "True"}::vars
-        val grob_z = [(rat_1,1::(map (fn x => 0) vars))]
-        val grob_1 = [(rat_1,(map (fn x => 0) vars'))]
+        val grob_z = [(rat_1,1::(map (K 0) vars))]
+        val grob_1 = [(rat_1,(map (K 0) vars'))]
         fun augment p= map (fn (c,m) => (c,0::m)) p
         val pols' = map augment pols
         val pol' = augment pol
@@ -387,7 +352,7 @@
 
 fun refute_disj rfn tm =
  case term_of tm of
-  Const(@{const_name HOL.disj},_)$l$r =>
+  Const(@{const_name HOL.disj},_)$_$_ =>
    Drule.compose
     (refute_disj rfn (Thm.dest_arg tm), 2,
       Drule.compose (refute_disj rfn (Thm.dest_arg1 tm), 2, disjE))
@@ -398,7 +363,7 @@
 
 fun is_neg t =
     case term_of t of
-      (Const(@{const_name Not},_)$p) => true
+      (Const(@{const_name Not},_)$_) => true
     | _  => false;
 fun is_eq t =
  case term_of t of
@@ -423,7 +388,7 @@
 val strip_exists =
  let fun h (acc, t) =
       case term_of t of
-       Const (@{const_name Ex}, _) $ Abs (x, T, p) =>
+       Const (@{const_name Ex}, _) $ Abs _ =>
         h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc))
      | _ => (acc,t)
  in fn t => h ([],t)
@@ -435,10 +400,7 @@
 | _ => false;
 
 val mk_object_eq = fn th => th COMP meta_eq_to_obj_eq;
-val bool_simps = @{thms bool_simps};
 val nnf_simps = @{thms nnf_simps};
-fun nnf_conv ctxt =
-  Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps bool_simps addsimps nnf_simps)
 
 fun weak_dnf_conv ctxt =
   Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms weak_dnf_simps});
@@ -484,12 +446,10 @@
 
 fun fold1 f = foldr1 (uncurry f);
 
-val list_conj = fold1 (fn c => fn c' => Thm.apply (Thm.apply @{cterm HOL.conj} c) c') ;
-
 fun mk_conj_tab th =
  let fun h acc th =
    case prop_of th of
-   @{term "Trueprop"}$(@{term HOL.conj}$p$q) =>
+   @{term "Trueprop"}$(@{term HOL.conj}$_$_) =>
      h (h acc (th RS conjunct2)) (th RS conjunct1)
   | @{term "Trueprop"}$p => (p,th)::acc
 in fold (Termtab.insert Thm.eq_thm) (h [] th) Termtab.empty end;
@@ -567,8 +527,7 @@
  | Var ((s,_),_) => s
  | _ => "x"
  fun mk_eq s t = Thm.apply (Thm.apply @{cterm "op == :: bool => _"} s) t
- fun mkeq s t = Thm.apply @{cterm Trueprop} (Thm.apply (Thm.apply @{cterm "op = :: bool => _"} s) t)
- fun mk_exists v th = Drule.arg_cong_rule (ext (ctyp_of_term v))
+  fun mk_exists v th = Drule.arg_cong_rule (ext (ctyp_of_term v))
    (Thm.abstract_rule (getname v) v th)
  fun simp_ex_conv ctxt =
    Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms(39)})
@@ -585,8 +544,8 @@
 (** main **)
 
 fun ring_and_ideal_conv
-  {vars, semiring = (sr_ops, sr_rules), ring = (r_ops, r_rules),
-   field = (f_ops, f_rules), idom, ideal}
+  {vars = _, semiring = (sr_ops, _), ring = (r_ops, _),
+   field = (f_ops, _), idom, ideal}
   dest_const mk_const ring_eq_conv ring_normalize_conv =
 let
   val [add_pat, mul_pat, pow_pat, zero_tm, one_tm] = sr_ops;
@@ -612,7 +571,6 @@
        else raise CTERM ("ring_dest_neg", [t])
     end
 
- val ring_mk_neg = fn tm => Thm.apply (ring_neg_tm) (tm);
  fun field_dest_inv t =
     let val (l,r) = Thm.dest_comb t in
         if Term.could_unify(term_of l, term_of field_inv_tm) then r
@@ -621,11 +579,9 @@
  val ring_dest_add = dest_binary ring_add_tm;
  val ring_mk_add = mk_binop ring_add_tm;
  val ring_dest_sub = dest_binary ring_sub_tm;
- val ring_mk_sub = mk_binop ring_sub_tm;
  val ring_dest_mul = dest_binary ring_mul_tm;
  val ring_mk_mul = mk_binop ring_mul_tm;
  val field_dest_div = dest_binary field_div_tm;
- val field_mk_div = mk_binop field_div_tm;
  val ring_dest_pow = dest_binary ring_pow_tm;
  val ring_mk_pow = mk_binop ring_pow_tm ;
  fun grobvars tm acc =
@@ -652,7 +608,7 @@
      [(rat_1,map (fn i => if i aconvc tm then 1 else 0) vars)])
 handle  CTERM _ =>
  ((let val x = dest_const tm
- in if x =/ rat_0 then [] else [(x,map (fn v => 0) vars)]
+ in if x =/ rat_0 then [] else [(x,map (K 0) vars)]
  end)
  handle ERROR _ =>
   ((grob_neg(grobify_term vars (ring_dest_neg tm)))
@@ -732,7 +688,7 @@
         Conv.fconv_rule
           ((Conv.arg_conv #> Conv.arg_conv) (Conv.binop_conv ring_normalize_conv)) th1
       val conc = th2 |> concl |> Thm.dest_arg
-      val (l,r) = conc |> dest_eq
+      val (l,_) = conc |> dest_eq
     in Thm.implies_intr (Thm.apply cTrp tm)
                     (Thm.equal_elim (Drule.arg_cong_rule cTrp (eqF_intr th2))
                            (Thm.reflexive l |> mk_object_eq))
@@ -756,9 +712,9 @@
        val th2 = funpow deg (idom_rule ctxt o HOLogic.conj_intr th1) neq_01
       in (vars,l,cert,th2)
       end)
-    val cert_pos = map (fn (i,p) => (i,filter (fn (c,m) => c >/ rat_0) p)) cert
+    val cert_pos = map (fn (i,p) => (i,filter (fn (c,_) => c >/ rat_0) p)) cert
     val cert_neg = map (fn (i,p) => (i,map (fn (c,m) => (minus_rat c,m))
-                                            (filter (fn (c,m) => c </ rat_0) p))) cert
+                                            (filter (fn (c,_) => c </ rat_0) p))) cert
     val  herts_pos = map (fn (i,p) => (i,holify_polynomial vars p)) cert_pos
     val  herts_neg = map (fn (i,p) => (i,holify_polynomial vars p)) cert_neg
     fun thm_fn pols =
@@ -772,7 +728,7 @@
     val th4 =
       Conv.fconv_rule ((Conv.arg_conv o Conv.arg_conv o Conv.binop_conv) ring_normalize_conv)
         (neq_rule l th3)
-    val (l,r) = dest_eq(Thm.dest_arg(concl th4))
+    val (l, _) = dest_eq(Thm.dest_arg(concl th4))
    in Thm.implies_intr (Thm.apply cTrp tm)
                         (Thm.equal_elim (Drule.arg_cong_rule cTrp (eqF_intr th4))
                    (Thm.reflexive l |> mk_object_eq))
@@ -873,7 +829,6 @@
       (Drule.binop_cong_rule @{cterm HOL.conj} th1
         (Thm.reflexive (Thm.dest_arg (Thm.rhs_of th2))))
   val v = Thm.dest_arg1(Thm.dest_arg1(Thm.rhs_of th3))
-  val vars' = (remove op aconvc v vars) @ [v]
   val th4 = Conv.fconv_rule (Conv.arg_conv (simp_ex_conv ctxt)) (mk_exists v th3)
   val th5 = ex_eq_conv (mk_eq tm (fold mk_ex (remove op aconvc v vars) (Thm.lhs_of th4)))
  in Thm.transitive th5 (fold mk_exists (remove op aconvc v vars) th4)
@@ -961,23 +916,12 @@
 | SOME tm =>
   (case Semiring_Normalizer.match ctxt tm of
     NONE => NONE
-  | SOME (res as (theory, {is_const, dest_const,
+  | SOME (res as (theory, {is_const = _, dest_const,
           mk_const, conv = ring_eq_conv})) =>
      SOME (ring_and_ideal_conv theory
           dest_const (mk_const (ctyp_of_term tm)) (ring_eq_conv ctxt)
           (Semiring_Normalizer.semiring_normalize_wrapper ctxt res)))
 
-fun ring_solve ctxt form =
-  (case try (find_term 0 (* FIXME !? *)) form of
-    NONE => Thm.reflexive form
-  | SOME tm =>
-      (case Semiring_Normalizer.match ctxt tm of
-        NONE => Thm.reflexive form
-      | SOME (res as (theory, {is_const, dest_const, mk_const, conv = ring_eq_conv})) =>
-        #ring_conv (ring_and_ideal_conv theory
-          dest_const (mk_const (ctyp_of_term tm)) (ring_eq_conv ctxt)
-          (Semiring_Normalizer.semiring_normalize_wrapper ctxt res)) ctxt form));
-
 fun presimplify ctxt add_thms del_thms =
   asm_full_simp_tac (put_simpset HOL_basic_ss ctxt
     addsimps (Algebra_Simplification.get ctxt)
@@ -1014,7 +958,7 @@
  | SOME thy =>
   let
    fun poly_exists_tac {asms = asms, concl = concl, prems = prems,
-            params = params, context = ctxt, schematics = scs} =
+            params = _, context = ctxt, schematics = _} =
     let
      val (evs,bod) = strip_exists (Thm.dest_arg concl)
      val ps = map_filter (try (lhs o Thm.dest_arg)) asms
--- a/src/HOL/Tools/group_cancel.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/group_cancel.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -25,7 +25,7 @@
 val sub1 = @{lemma "(A::'a::ab_group_add) == k + a ==> A - b == k + (a - b)"
       by (simp only: add_diff_eq)}
 val sub2 = @{lemma "(B::'a::ab_group_add) == k + b ==> a - B == - k + (a - b)"
-      by (simp only: diff_minus minus_add add_ac)}
+      by (simp only: minus_add diff_conv_add_uminus add_ac)}
 val neg1 = @{lemma "(A::'a::ab_group_add) == k + a ==> - A == - k + - a"
       by (simp only: minus_add_distrib)}
 val rule0 = @{lemma "(a::'a::comm_monoid_add) == a + 0"
--- a/src/HOL/Tools/int_arith.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/int_arith.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -87,9 +87,9 @@
 val setup =
   Lin_Arith.add_inj_thms [@{thm zle_int} RS iffD2, @{thm int_int_eq} RS iffD2]
   #> Lin_Arith.add_lessD @{thm zless_imp_add1_zle}
-  #> Lin_Arith.add_simps (@{thms simp_thms} @ @{thms arith_simps} @ @{thms rel_simps}
-      @ @{thms pred_numeral_simps}
-      @ @{thms arith_special} @ @{thms int_arith_rules})
+  #> Lin_Arith.add_simps @{thms of_nat_simps of_int_simps}
+  #> Lin_Arith.add_simps
+      [@{thm of_int_numeral}, @{thm nat_0}, @{thm nat_1}, @{thm diff_nat_numeral}, @{thm nat_numeral}]
   #> Lin_Arith.add_simprocs [zero_one_idom_simproc]
   #> Lin_Arith.set_number_of number_of
   #> Lin_Arith.add_inj_const (@{const_name of_nat}, HOLogic.natT --> HOLogic.intT)
--- a/src/HOL/Tools/lin_arith.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/lin_arith.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -791,37 +791,16 @@
    Most of the work is done by the cancel tactics. *)
 
 val init_arith_data =
-  Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, number_of, ...} =>
-   {add_mono_thms = @{thms add_mono_thms_linordered_semiring} @
-      @{thms add_mono_thms_linordered_field} @ add_mono_thms,
-    mult_mono_thms = @{thm mult_strict_left_mono} :: @{thm mult_left_mono} ::
-      @{lemma "a = b ==> c*a = c*b" by (rule arg_cong)} :: mult_mono_thms,
+  Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, number_of, ...} =>
+   {add_mono_thms = @{thms add_mono_thms_linordered_semiring}
+      @ @{thms add_mono_thms_linordered_field} @ add_mono_thms,
+    mult_mono_thms = @{thm mult_strict_left_mono} :: @{thm mult_left_mono}
+      :: @{lemma "a = b ==> c * a = c * b" by (rule arg_cong)} :: mult_mono_thms,
     inj_thms = inj_thms,
-    lessD = lessD @ [@{thm "Suc_leI"}],
-    neqE = [@{thm linorder_neqE_nat}, @{thm linorder_neqE_linordered_idom}],
-    simpset =
-      put_simpset HOL_basic_ss @{context}
-      addsimps @{thms ring_distribs}
-      addsimps [@{thm if_True}, @{thm if_False}]
-      addsimps
-       [@{thm add_0_left}, @{thm add_0_right},
-        @{thm add_Suc}, @{thm add_Suc_right},
-        @{thm nat.inject}, @{thm Suc_le_mono}, @{thm Suc_less_eq},
-        @{thm "Zero_not_Suc"}, @{thm "Suc_not_Zero"}, @{thm "le_0_eq"}, @{thm "One_nat_def"},
-        @{thm "order_less_irrefl"}, @{thm "zero_neq_one"}, @{thm "zero_less_one"},
-        @{thm "zero_le_one"}, @{thm "zero_neq_one"} RS not_sym, @{thm "not_one_le_zero"},
-        @{thm "not_one_less_zero"}]
-      addsimprocs [@{simproc group_cancel_add}, @{simproc group_cancel_diff},
-                   @{simproc group_cancel_eq}, @{simproc group_cancel_le},
-                   @{simproc group_cancel_less}]
-       (*abel_cancel helps it work in abstract algebraic domains*)
-      addsimprocs [@{simproc nateq_cancel_sums},
-                   @{simproc natless_cancel_sums},
-                   @{simproc natle_cancel_sums}]
-      |> Simplifier.add_cong @{thm if_weak_cong}
-      |> simpset_of,
-    number_of = number_of}) #>
-  add_discrete_type @{type_name nat};
+    lessD = lessD,
+    neqE = @{thm linorder_neqE_nat} :: @{thm linorder_neqE_linordered_idom} :: neqE,
+    simpset = put_simpset HOL_basic_ss @{context} |> Simplifier.add_cong @{thm if_weak_cong} |> simpset_of,
+    number_of = number_of});
 
 (* FIXME !?? *)
 fun add_arith_facts ctxt =
@@ -909,9 +888,6 @@
 
 (* context setup *)
 
-val setup =
-  init_arith_data;
-
 val global_setup =
   map_theory_simpset (fn ctxt => ctxt
     addSolver (mk_solver "lin_arith" (add_arith_facts #> Fast_Arith.prems_lin_arith_tac))) #>
@@ -924,4 +900,22 @@
           THEN' tac ctxt)))) "linear arithmetic" #>
   Arith_Data.add_tactic "linear arithmetic" gen_tac;
 
+val setup =
+  init_arith_data
+  #> add_discrete_type @{type_name nat}
+  #> add_lessD @{thm Suc_leI}
+  #> add_simps (@{thms simp_thms} @ @{thms ring_distribs} @ [@{thm if_True}, @{thm if_False},
+      @{thm add_0_left}, @{thm add_0_right}, @{thm order_less_irrefl},
+      @{thm zero_neq_one}, @{thm zero_less_one}, @{thm zero_le_one},
+      @{thm zero_neq_one} RS not_sym, @{thm not_one_le_zero}, @{thm not_one_less_zero}])
+  #> add_simps [@{thm add_Suc}, @{thm add_Suc_right}, @{thm nat.inject},
+      @{thm Suc_le_mono}, @{thm Suc_less_eq}, @{thm Zero_not_Suc},
+      @{thm Suc_not_Zero}, @{thm le_0_eq}, @{thm One_nat_def}]
+  #> add_simprocs [@{simproc group_cancel_add}, @{simproc group_cancel_diff},
+      @{simproc group_cancel_eq}, @{simproc group_cancel_le},
+      @{simproc group_cancel_less}]
+     (*abel_cancel helps it work in abstract algebraic domains*)
+  #> add_simprocs [@{simproc nateq_cancel_sums},@{simproc natless_cancel_sums},
+      @{simproc natle_cancel_sums}];
+
 end;
--- a/src/HOL/Tools/numeral_simprocs.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/numeral_simprocs.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -220,7 +220,7 @@
 val minus_simps = [@{thm minus_zero}, @{thm minus_one}, @{thm minus_numeral}, @{thm minus_neg_numeral}];
 
 (*To let us treat subtraction as addition*)
-val diff_simps = [@{thm diff_minus}, @{thm minus_add_distrib}, @{thm minus_minus}];
+val diff_simps = [@{thm diff_conv_add_uminus}, @{thm minus_add_distrib}, @{thm minus_minus}];
 
 (*To let us treat division as multiplication*)
 val divide_simps = [@{thm divide_inverse}, @{thm inverse_mult_distrib}, @{thm inverse_inverse_eq}];
@@ -719,7 +719,7 @@
            @{thm "times_divide_eq_left"}, @{thm "times_divide_eq_right"},
            @{thm "times_divide_times_eq"},
            @{thm "divide_divide_eq_right"},
-           @{thm "diff_minus"}, @{thm "minus_divide_left"},
+           @{thm diff_conv_add_uminus}, @{thm "minus_divide_left"},
            @{thm "add_divide_distrib"} RS sym,
            @{thm field_divide_inverse} RS sym, @{thm inverse_divide}, 
            Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Conv.rewr_conv (mk_meta_eq @{thm mult_commute}))))   
--- a/src/HOL/Tools/semiring_normalizer.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/semiring_normalizer.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -848,7 +848,7 @@
 val nat_exp_ss =
   simpset_of
    (put_simpset HOL_basic_ss @{context}
-    addsimps (@{thms eval_nat_numeral} @ @{thms nat_arith} @ @{thms arith_simps} @ @{thms rel_simps})
+    addsimps (@{thms eval_nat_numeral} @ @{thms diff_nat_numeral} @ @{thms arith_simps} @ @{thms rel_simps})
     addsimps [@{thm Let_def}, @{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc}]);
 
 fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS;
--- a/src/HOL/Tools/try0.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Tools/try0.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -41,10 +41,10 @@
   end
   handle TimeLimit.TimeOut => false
 
-fun do_generic timeout_opt command pre post apply st =
+fun do_generic timeout_opt name command pre post apply st =
   let val timer = Timer.startRealTimer () in
     if can_apply timeout_opt pre post apply st then
-      SOME (command, Time.toMilliseconds (Timer.checkRealTimer timer))
+      SOME (name, command, Time.toMilliseconds (Timer.checkRealTimer timer))
     else
       NONE
   end
@@ -75,16 +75,11 @@
                     timeout_opt quad st =
   if mode <> Auto_Try orelse run_if_auto_try then
     let val attrs = attrs_text attrs quad in
-      do_generic timeout_opt
-                 (name ^ attrs ^
-                  (if all_goals andalso
-                      nprems_of (#goal (Proof.goal st)) > 1 then
-                     " [1]"
-                   else
-                     ""))
-                 I (#goal o Proof.goal)
-                 (apply_named_method_on_first_goal (name ^ attrs)
-                                                   (Proof.theory_of st)) st
+      do_generic timeout_opt name
+        ((name ^ attrs |> attrs <> "" ? enclose "(" ")") ^
+         (if all_goals andalso nprems_of (#goal (Proof.goal st)) > 1 then "[1]" else ""))
+        I (#goal o Proof.goal)
+        (apply_named_method_on_first_goal (name ^ attrs) (Proof.theory_of st)) st
     end
   else
     NONE
@@ -108,12 +103,16 @@
    ("presburger", ((false, true), no_attrs))]
 val do_methods = map do_named_method named_methods
 
-fun time_string (s, ms) = s ^ ": " ^ string_of_int ms ^ " ms"
+fun time_string ms = string_of_int ms ^ " ms"
+fun tool_time_string (s, ms) = s ^ ": " ^ time_string ms
 
 fun do_try0 mode timeout_opt quad st =
   let
     val st = st |> Proof.map_context (Config.put Metis_Tactic.verbose false #>
       Config.put Lin_Arith.verbose false)
+    fun par_map f =
+      if mode = Normal then Par_List.map f #> map_filter I #> sort (int_ord o pairself #3)
+      else Par_List.get_some f #> the_list
   in
     if mode = Normal then
       "Trying " ^ space_implode " " (Try.serial_commas "and"
@@ -121,17 +120,15 @@
       |> Output.urgent_message
     else
       ();
-    case do_methods |> Par_List.map (fn f => f mode timeout_opt quad st)
-                    |> map_filter I |> sort (int_ord o pairself snd) of
+    (case par_map (fn f => f mode timeout_opt quad st) do_methods of
       [] =>
       (if mode = Normal then Output.urgent_message "No proof found." else ();
        (false, (noneN, st)))
-    | xs as (s, _) :: _ =>
+    | xs as (name, command, _) :: _ =>
       let
-        val xs = xs |> map (fn (s, n) => (n, hd (space_explode " " s)))
+        val xs = xs |> map (fn (name, _, n) => (n, name))
                     |> AList.coalesce (op =)
                     |> map (swap o apsnd commas)
-        val need_parens = exists_string (curry (op =) " ") s
         val message =
           (case mode of
              Auto_Try => "Auto Try0 found a proof"
@@ -139,16 +136,18 @@
            | Normal => "Try this") ^ ": " ^
           Active.sendback_markup [Markup.padding_command]
               ((if nprems_of (#goal (Proof.goal st)) = 1 then "by"
-                else "apply") ^ " " ^ (s |> need_parens ? enclose "(" ")")) ^
-          "\n(" ^ space_implode "; " (map time_string xs) ^ ")."
+                else "apply") ^ " " ^ command) ^
+          (case xs of
+            [(_, ms)] => " (" ^ time_string ms ^ ")."
+          | xs => "\n(" ^ space_implode "; " (map tool_time_string xs) ^ ").")
       in
-        (true, (s, st |> (if mode = Auto_Try then
-                            Proof.goal_message
-                                (fn () => Pretty.markup Markup.information
-                                                        [Pretty.str message])
-                          else
-                            tap (fn _ => Output.urgent_message message))))
-      end
+        (true, (name,
+           st |> (if mode = Auto_Try then
+                    Proof.goal_message
+                      (fn () => Pretty.markup Markup.information [Pretty.str message])
+                  else
+                    tap (fn _ => Output.urgent_message message))))
+      end)
   end
 
 fun try0 timeout_opt = fst oo do_try0 Normal timeout_opt
--- a/src/HOL/Topological_Spaces.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Topological_Spaces.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -2112,7 +2112,7 @@
   with dense[of b "Inf A"] obtain c where "c < Inf A" "c \<in> A"
     by (auto simp: subset_eq)
   then show False
-    using cInf_lower[OF `c \<in> A`, of x] bnd by (metis less_imp_le not_le)
+    using cInf_lower[OF `c \<in> A`] bnd by (metis not_le less_imp_le bdd_belowI)
 qed
 
 lemma Sup_notin_open:
@@ -2125,7 +2125,7 @@
   with dense[of "Sup A" b] obtain c where "Sup A < c" "c \<in> A"
     by (auto simp: subset_eq)
   then show False
-    using cSup_upper[OF `c \<in> A`, of x] bnd by (metis less_imp_le not_le)
+    using cSup_upper[OF `c \<in> A`] bnd by (metis less_imp_le not_le bdd_aboveI)
 qed
 
 end
@@ -2151,7 +2151,7 @@
     let ?z = "Inf (B \<inter> {x <..})"
 
     have "x \<le> ?z" "?z \<le> y"
-      using `y \<in> B` `x < y` by (auto intro: cInf_lower[where z=x] cInf_greatest)
+      using `y \<in> B` `x < y` by (auto intro: cInf_lower cInf_greatest)
     with `x \<in> U` `y \<in> U` have "?z \<in> U"
       by (rule *)
     moreover have "?z \<notin> B \<inter> {x <..}"
@@ -2163,11 +2163,11 @@
       obtain a where "?z < a" "{?z ..< a} \<subseteq> A"
         using open_right[OF `open A` `?z \<in> A` `?z < y`] by auto
       moreover obtain b where "b \<in> B" "x < b" "b < min a y"
-        using cInf_less_iff[of "B \<inter> {x <..}" x "min a y"] `?z < a` `?z < y` `x < y` `y \<in> B`
+        using cInf_less_iff[of "B \<inter> {x <..}" "min a y"] `?z < a` `?z < y` `x < y` `y \<in> B`
         by (auto intro: less_imp_le)
       moreover have "?z \<le> b"
         using `b \<in> B` `x < b`
-        by (intro cInf_lower[where z=x]) auto
+        by (intro cInf_lower) auto
       moreover have "b \<in> U"
         using `x \<le> ?z` `?z \<le> b` `b < min a y`
         by (intro *[OF `x \<in> U` `y \<in> U`]) (auto simp: less_imp_le)
--- a/src/HOL/Transcendental.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Transcendental.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -453,7 +453,7 @@
   apply simp
   apply (simp only: lemma_termdiff1 setsum_right_distrib)
   apply (rule setsum_cong [OF refl])
-  apply (simp add: diff_minus [symmetric] less_iff_Suc_add)
+  apply (simp add: less_iff_Suc_add)
   apply (clarify)
   apply (simp add: setsum_right_distrib lemma_realpow_diff_sumr2 mult_ac
               del: setsum_op_ivl_Suc power_Suc)
@@ -1129,8 +1129,7 @@
   by (rule inverse_unique [symmetric], simp add: mult_exp_exp)
 
 lemma exp_diff: "exp (x - y) = exp x / exp y"
-  unfolding diff_minus divide_inverse
-  by (simp add: exp_add exp_minus)
+  using exp_add [of x "- y"] by (simp add: exp_minus divide_inverse)
 
 
 subsubsection {* Properties of the Exponential Function on Reals *}
@@ -2410,13 +2409,13 @@
   using sin_cos_minus_lemma [where x=x] by simp
 
 lemma sin_diff: "sin (x - y) = sin x * cos y - cos x * sin y"
-  by (simp add: diff_minus sin_add)
+  using sin_add [of x "- y"] by simp
 
 lemma sin_diff2: "sin (x - y) = cos y * sin x - sin y * cos x"
   by (simp add: sin_diff mult_commute)
 
 lemma cos_diff: "cos (x - y) = cos x * cos y + sin x * sin y"
-  by (simp add: diff_minus cos_add)
+  using cos_add [of x "- y"] by simp
 
 lemma cos_diff2: "cos (x - y) = cos y * cos x + sin y * sin x"
   by (simp add: cos_diff mult_commute)
@@ -2526,8 +2525,9 @@
         by (simp add: inverse_eq_divide less_divide_eq)
     }
     note *** = this
+    have [simp]: "\<And>x y::real. 0 < x - y \<longleftrightarrow> y < x" by arith
     from ** show ?thesis by (rule sumr_pos_lt_pair)
-      (simp add: divide_inverse real_0_less_add_iff mult_assoc [symmetric] ***)
+      (simp add: divide_inverse mult_assoc [symmetric] ***)
   qed
   ultimately have "0 < (\<Sum>n. - (-1 ^ n * 2 ^ (2 * n) / real (fact (2 * n))))"
     by (rule order_less_trans)
@@ -2810,7 +2810,7 @@
 apply (cut_tac x="-x" in cos_zero_lemma, simp, simp)
 apply (force simp add: minus_equation_iff [of x])
 apply (auto simp only: odd_Suc_mult_two_ex real_of_nat_Suc distrib_right)
-apply (auto simp add: cos_add)
+apply (auto simp add: cos_diff cos_add)
 done
 
 (* ditto: but to a lesser extent *)
@@ -3833,7 +3833,7 @@
               by (rule DERIV_arctan_suminf[OF `0 < r` `r < 1` `\<bar>x\<bar> < r`])
             from DERIV_add_minus[OF this DERIV_arctan]
             show "DERIV (\<lambda> x. suminf (?c x) - arctan x) x :> 0"
-              unfolding diff_minus by auto
+              by auto
           qed
           hence DERIV_in_rball: "\<forall> y. a \<le> y \<and> y \<le> b \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) y :> 0"
             using `-r < a` `b < r` by auto
@@ -3922,9 +3922,10 @@
       }
       hence "\<forall> x \<in> { 0 <..< 1 }. 0 \<le> ?a x n - ?diff x n" by auto
       moreover have "\<And>x. isCont (\<lambda> x. ?a x n - ?diff x n) x"
-        unfolding diff_minus divide_inverse
+        unfolding diff_conv_add_uminus divide_inverse
         by (auto intro!: isCont_add isCont_rabs isCont_ident isCont_minus isCont_arctan
-          isCont_inverse isCont_mult isCont_power isCont_const isCont_setsum)
+          isCont_inverse isCont_mult isCont_power isCont_const isCont_setsum
+          simp del: add_uminus_conv_diff)
       ultimately have "0 \<le> ?a 1 n - ?diff 1 n"
         by (rule LIM_less_bound)
       hence "?diff 1 n \<le> ?a 1 n" by auto
--- a/src/HOL/Wellfounded.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Wellfounded.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -482,6 +482,11 @@
 
 lemmas accpI = accp.accI
 
+lemma accp_eq_acc [code]:
+  "accp r = (\<lambda>x. x \<in> Wellfounded.acc {(x, y). r x y})"
+  by (simp add: acc_def)
+
+
 text {* Induction rules *}
 
 theorem accp_induct:
@@ -855,4 +860,7 @@
 
 declare "prod.size" [no_atp]
 
+
+hide_const (open) acc accp
+
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Word/Bit_Bit.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -0,0 +1,73 @@
+(*  Title:      HOL/Word/Bit_Bit.thy
+    Author:     Author: Brian Huffman, PSU and Gerwin Klein, NICTA
+*)
+
+header {* Bit operations in \<Z>\<^sub>2*}
+
+theory Bit_Bit
+imports Bit_Operations "~~/src/HOL/Library/Bit"
+begin
+
+instantiation bit :: bit
+begin
+
+primrec bitNOT_bit where
+  "NOT 0 = (1::bit)"
+  | "NOT 1 = (0::bit)"
+
+primrec bitAND_bit where
+  "0 AND y = (0::bit)"
+  | "1 AND y = (y::bit)"
+
+primrec bitOR_bit where
+  "0 OR y = (y::bit)"
+  | "1 OR y = (1::bit)"
+
+primrec bitXOR_bit where
+  "0 XOR y = (y::bit)"
+  | "1 XOR y = (NOT y :: bit)"
+
+instance  ..
+
+end
+
+lemmas bit_simps =
+  bitNOT_bit.simps bitAND_bit.simps bitOR_bit.simps bitXOR_bit.simps
+
+lemma bit_extra_simps [simp]: 
+  "x AND 0 = (0::bit)"
+  "x AND 1 = (x::bit)"
+  "x OR 1 = (1::bit)"
+  "x OR 0 = (x::bit)"
+  "x XOR 1 = NOT (x::bit)"
+  "x XOR 0 = (x::bit)"
+  by (cases x, auto)+
+
+lemma bit_ops_comm: 
+  "(x::bit) AND y = y AND x"
+  "(x::bit) OR y = y OR x"
+  "(x::bit) XOR y = y XOR x"
+  by (cases y, auto)+
+
+lemma bit_ops_same [simp]: 
+  "(x::bit) AND x = x"
+  "(x::bit) OR x = x"
+  "(x::bit) XOR x = 0"
+  by (cases x, auto)+
+
+lemma bit_not_not [simp]: "NOT (NOT (x::bit)) = x"
+  by (cases x) auto
+
+lemma bit_or_def: "(b::bit) OR c = NOT (NOT b AND NOT c)"
+  by (induct b, simp_all)
+
+lemma bit_xor_def: "(b::bit) XOR c = (b AND NOT c) OR (NOT b AND c)"
+  by (induct b, simp_all)
+
+lemma bit_NOT_eq_1_iff [simp]: "NOT (b::bit) = 1 \<longleftrightarrow> b = 0"
+  by (induct b, simp_all)
+
+lemma bit_AND_eq_1_iff [simp]: "(a::bit) AND b = 1 \<longleftrightarrow> a = 1 \<and> b = 1"
+  by (induct a, simp_all)
+
+end
--- a/src/HOL/Word/Bit_Int.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Word/Bit_Int.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -9,7 +9,7 @@
 header {* Bitwise Operations on Binary Integers *}
 
 theory Bit_Int
-imports Bit_Representation Bit_Operations
+imports Bit_Representation Bit_Bit
 begin
 
 subsection {* Logical operations *}
--- a/src/HOL/Word/Bit_Operations.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Word/Bit_Operations.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -8,8 +8,6 @@
 imports "~~/src/HOL/Library/Bit"
 begin
 
-subsection {* Abstract syntactic bit operations *}
-
 class bit =
   fixes bitNOT :: "'a \<Rightarrow> 'a"       ("NOT _" [70] 71)
     and bitAND :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "AND" 64)
@@ -37,69 +35,5 @@
 class bitss = bits +
   fixes msb      :: "'a \<Rightarrow> bool"
 
-  
-subsection {* Bitwise operations on @{typ bit} *}
-
-instantiation bit :: bit
-begin
-
-primrec bitNOT_bit where
-  "NOT 0 = (1::bit)"
-  | "NOT 1 = (0::bit)"
-
-primrec bitAND_bit where
-  "0 AND y = (0::bit)"
-  | "1 AND y = (y::bit)"
-
-primrec bitOR_bit where
-  "0 OR y = (y::bit)"
-  | "1 OR y = (1::bit)"
-
-primrec bitXOR_bit where
-  "0 XOR y = (y::bit)"
-  | "1 XOR y = (NOT y :: bit)"
-
-instance  ..
-
 end
 
-lemmas bit_simps =
-  bitNOT_bit.simps bitAND_bit.simps bitOR_bit.simps bitXOR_bit.simps
-
-lemma bit_extra_simps [simp]: 
-  "x AND 0 = (0::bit)"
-  "x AND 1 = (x::bit)"
-  "x OR 1 = (1::bit)"
-  "x OR 0 = (x::bit)"
-  "x XOR 1 = NOT (x::bit)"
-  "x XOR 0 = (x::bit)"
-  by (cases x, auto)+
-
-lemma bit_ops_comm: 
-  "(x::bit) AND y = y AND x"
-  "(x::bit) OR y = y OR x"
-  "(x::bit) XOR y = y XOR x"
-  by (cases y, auto)+
-
-lemma bit_ops_same [simp]: 
-  "(x::bit) AND x = x"
-  "(x::bit) OR x = x"
-  "(x::bit) XOR x = 0"
-  by (cases x, auto)+
-
-lemma bit_not_not [simp]: "NOT (NOT (x::bit)) = x"
-  by (cases x) auto
-
-lemma bit_or_def: "(b::bit) OR c = NOT (NOT b AND NOT c)"
-  by (induct b, simp_all)
-
-lemma bit_xor_def: "(b::bit) XOR c = (b AND NOT c) OR (NOT b AND c)"
-  by (induct b, simp_all)
-
-lemma bit_NOT_eq_1_iff [simp]: "NOT (b::bit) = 1 \<longleftrightarrow> b = 0"
-  by (induct b, simp_all)
-
-lemma bit_AND_eq_1_iff [simp]: "(a::bit) AND b = 1 \<longleftrightarrow> a = 1 \<and> b = 1"
-  by (induct a, simp_all)
-
-end
--- a/src/HOL/Word/Bit_Representation.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Word/Bit_Representation.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -636,12 +636,12 @@
   unfolding no_sbintr_alt2 by (drule sb_inc_lem') simp
 
 lemma sb_dec_lem:
-  "(0::int) <= - (2^k) + a ==> (a + 2^k) mod (2 * 2 ^ k) <= - (2 ^ k) + a"
-  by (rule int_mod_le' [where n = "2 ^ (Suc k)" and b = "a + 2 ^ k", simplified])
+  "(0::int) \<le> - (2 ^ k) + a \<Longrightarrow> (a + 2 ^ k) mod (2 * 2 ^ k) \<le> - (2 ^ k) + a"
+  using int_mod_le'[where n = "2 ^ (Suc k)" and b = "a + 2 ^ k"] by simp
 
 lemma sb_dec_lem':
-  "(2::int) ^ k <= a ==> (a + 2 ^ k) mod (2 * 2 ^ k) <= - (2 ^ k) + a"
-  by (rule iffD1 [OF diff_le_eq', THEN sb_dec_lem, simplified])
+  "(2::int) ^ k \<le> a \<Longrightarrow> (a + 2 ^ k) mod (2 * 2 ^ k) \<le> - (2 ^ k) + a"
+  by (rule sb_dec_lem) simp
 
 lemma sbintrunc_dec:
   "x >= (2 ^ n) ==> x - 2 ^ (Suc n) >= sbintrunc n x"
--- a/src/HOL/Word/Misc_Numeric.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Word/Misc_Numeric.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -8,10 +8,6 @@
 imports Main Parity
 begin
 
-lemma zmod_zsub_self [simp]: (* FIXME move to Divides.thy *) 
-  "((b :: int) - a) mod a = b mod a"
-  by (simp add: mod_diff_right_eq)
-
 declare iszero_0 [iff]
 
 lemma min_pm [simp]: "min a b + (a - b) = (a :: nat)" by arith
--- a/src/HOL/Word/Word.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Word/Word.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -8,6 +8,7 @@
 imports
   Type_Length
   "~~/src/HOL/Library/Boolean_Algebra"
+  Bit_Bit
   Bool_List_Representation
   Misc_Typedef
   Word_Miscellaneous
@@ -505,10 +506,6 @@
 definition max_word :: "'a::len word" -- "Largest representable machine integer." where
   "max_word = word_of_int (2 ^ len_of TYPE('a) - 1)"
 
-primrec of_bool :: "bool \<Rightarrow> 'a::len word" where
-  "of_bool False = 0"
-| "of_bool True = 1"
-
 (* FIXME: only provide one theorem name *)
 lemmas of_nth_def = word_set_bits_def
 
@@ -4238,7 +4235,7 @@
 
 lemma max_word_max [simp,intro!]: "n \<le> max_word"
   by (cases n rule: word_int_cases)
-     (simp add: max_word_def word_le_def int_word_uint int_mod_eq')
+    (simp add: max_word_def word_le_def int_word_uint int_mod_eq' del: minus_mod_self1)
   
 lemma word_of_int_2p_len: "word_of_int (2 ^ len_of TYPE('a)) = (0::'a::len0 word)"
   by (subst word_uint.Abs_norm [symmetric]) simp
--- a/src/HOL/Word/WordBitwise.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/Word/WordBitwise.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -65,7 +65,7 @@
 
 lemma bl_word_sub:
   "to_bl (x - y) = to_bl (x + (- y))"
-  by (simp add: diff_def)
+  by simp
 
 lemma rbl_word_1:
   "rev (to_bl (1 :: ('a :: len0) word))
--- a/src/HOL/ex/Dedekind_Real.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/ex/Dedekind_Real.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -1506,7 +1506,6 @@
 instance real :: linorder
   by (intro_classes, rule real_le_linear)
 
-
 lemma real_le_eq_diff: "(x \<le> y) = (x-y \<le> (0::real))"
 apply (cases x, cases y) 
 apply (auto simp add: real_le real_zero_def real_diff_def real_add real_minus
@@ -1520,14 +1519,14 @@
   have "z + x - (z + y) = (z + -z) + (x - y)" 
     by (simp add: algebra_simps) 
   with le show ?thesis 
-    by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"] diff_minus)
+    by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"])
 qed
 
 lemma real_sum_gt_zero_less: "(0 < S + (-W::real)) ==> (W < S)"
-by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus)
+by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S])
 
 lemma real_less_sum_gt_zero: "(W < S) ==> (0 < S + (-W::real))"
-by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus)
+by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S])
 
 lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y"
 apply (cases x, cases y)
@@ -1543,7 +1542,7 @@
 apply (rule real_sum_gt_zero_less)
 apply (drule real_less_sum_gt_zero [of x y])
 apply (drule real_mult_order, assumption)
-apply (simp add: distrib_left)
+apply (simp add: algebra_simps)
 done
 
 instantiation real :: distrib_lattice
@@ -1657,7 +1656,6 @@
 lemma real_less_all_real2: "~ 0 < y ==> \<forall>x. y < real_of_preal x"
 by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1])
 
-
 subsection {* Completeness of Positive Reals *}
 
 text {*
@@ -1759,107 +1757,23 @@
 qed
 
 text {*
-  \medskip Completeness properties using @{text "isUb"}, @{text "isLub"} etc.
-*}
-
-lemma posreals_complete:
-  assumes positive_S: "\<forall>x \<in> S. 0 < x"
-    and not_empty_S: "\<exists>x. x \<in> S"
-    and upper_bound_Ex: "\<exists>u. isUb (UNIV::real set) S u"
-  shows "\<exists>t. isLub (UNIV::real set) S t"
-proof
-  let ?pS = "{w. real_of_preal w \<in> S}"
-
-  obtain u where "isUb UNIV S u" using upper_bound_Ex ..
-  hence sup: "\<forall>x \<in> S. x \<le> u" by (simp add: isUb_def setle_def)
-
-  obtain x where x_in_S: "x \<in> S" using not_empty_S ..
-  hence x_gt_zero: "0 < x" using positive_S by simp
-  have  "x \<le> u" using sup and x_in_S ..
-  hence "0 < u" using x_gt_zero by arith
-
-  then obtain pu where u_is_pu: "u = real_of_preal pu"
-    by (auto simp add: real_gt_zero_preal_Ex)
-
-  have pS_less_pu: "\<forall>pa \<in> ?pS. pa \<le> pu"
-  proof
-    fix pa
-    assume "pa \<in> ?pS"
-    then obtain a where a: "a \<in> S" "a = real_of_preal pa"
-      by simp
-    then have "a \<le> u" using sup by simp
-    with a show "pa \<le> pu"
-      using sup and u_is_pu by (simp add: real_of_preal_le_iff)
-  qed
-
-  have "\<forall>y \<in> S. y \<le> real_of_preal (psup ?pS)"
-  proof
-    fix y
-    assume y_in_S: "y \<in> S"
-    hence "0 < y" using positive_S by simp
-    then obtain py where y_is_py: "y = real_of_preal py"
-      by (auto simp add: real_gt_zero_preal_Ex)
-    hence py_in_pS: "py \<in> ?pS" using y_in_S by simp
-    with pS_less_pu have "py \<le> psup ?pS"
-      by (rule preal_psup_le)
-    thus "y \<le> real_of_preal (psup ?pS)"
-      using y_is_py by (simp add: real_of_preal_le_iff)
-  qed
-
-  moreover {
-    fix x
-    assume x_ub_S: "\<forall>y\<in>S. y \<le> x"
-    have "real_of_preal (psup ?pS) \<le> x"
-    proof -
-      obtain "s" where s_in_S: "s \<in> S" using not_empty_S ..
-      hence s_pos: "0 < s" using positive_S by simp
-
-      hence "\<exists> ps. s = real_of_preal ps" by (simp add: real_gt_zero_preal_Ex)
-      then obtain "ps" where s_is_ps: "s = real_of_preal ps" ..
-      hence ps_in_pS: "ps \<in> {w. real_of_preal w \<in> S}" using s_in_S by simp
-
-      from x_ub_S have "s \<le> x" using s_in_S ..
-      hence "0 < x" using s_pos by simp
-      hence "\<exists> px. x = real_of_preal px" by (simp add: real_gt_zero_preal_Ex)
-      then obtain "px" where x_is_px: "x = real_of_preal px" ..
-
-      have "\<forall>pe \<in> ?pS. pe \<le> px"
-      proof
-        fix pe
-        assume "pe \<in> ?pS"
-        hence "real_of_preal pe \<in> S" by simp
-        hence "real_of_preal pe \<le> x" using x_ub_S by simp
-        thus "pe \<le> px" using x_is_px by (simp add: real_of_preal_le_iff)
-      qed
-
-      moreover have "?pS \<noteq> {}" using ps_in_pS by auto
-      ultimately have "(psup ?pS) \<le> px" by (simp add: psup_le_ub)
-      thus "real_of_preal (psup ?pS) \<le> x" using x_is_px by (simp add: real_of_preal_le_iff)
-    qed
-  }
-  ultimately show "isLub UNIV S (real_of_preal (psup ?pS))"
-    by (simp add: isLub_def leastP_def isUb_def setle_def setge_def)
-qed
-
-text {*
-  \medskip reals Completeness (again!)
+  \medskip Completeness
 *}
 
 lemma reals_complete:
+  fixes S :: "real set"
   assumes notempty_S: "\<exists>X. X \<in> S"
-    and exists_Ub: "\<exists>Y. isUb (UNIV::real set) S Y"
-  shows "\<exists>t. isLub (UNIV :: real set) S t"
+    and exists_Ub: "bdd_above S"
+  shows "\<exists>x. (\<forall>s\<in>S. s \<le> x) \<and> (\<forall>y. (\<forall>s\<in>S. s \<le> y) \<longrightarrow> x \<le> y)"
 proof -
   obtain X where X_in_S: "X \<in> S" using notempty_S ..
-  obtain Y where Y_isUb: "isUb (UNIV::real set) S Y"
-    using exists_Ub ..
+  obtain Y where Y_isUb: "\<forall>s\<in>S. s \<le> Y"
+    using exists_Ub by (auto simp: bdd_above_def)
   let ?SHIFT = "{z. \<exists>x \<in>S. z = x + (-X) + 1} \<inter> {x. 0 < x}"
 
   {
     fix x
-    assume "isUb (UNIV::real set) S x"
-    hence S_le_x: "\<forall> y \<in> S. y <= x"
-      by (simp add: isUb_def setle_def)
+    assume S_le_x: "\<forall>s\<in>S. s \<le> x"
     {
       fix s
       assume "s \<in> {z. \<exists>x\<in>S. z = x + - X + 1}"
@@ -1868,86 +1782,74 @@
       then have "x1 \<le> x" using S_le_x by simp
       with x1 have "s \<le> x + - X + 1" by arith
     }
-    then have "isUb (UNIV::real set) ?SHIFT (x + (-X) + 1)"
-      by (auto simp add: isUb_def setle_def)
+    then have "\<forall>s\<in>?SHIFT. s \<le> x + (-X) + 1"
+      by auto
   } note S_Ub_is_SHIFT_Ub = this
 
-  hence "isUb UNIV ?SHIFT (Y + (-X) + 1)" using Y_isUb by simp
-  hence "\<exists>Z. isUb UNIV ?SHIFT Z" ..
+  have *: "\<forall>s\<in>?SHIFT. s \<le> Y + (-X) + 1" using Y_isUb by (rule S_Ub_is_SHIFT_Ub)
+  have "\<forall>s\<in>?SHIFT. s < Y + (-X) + 2"
+  proof
+    fix s assume "s\<in>?SHIFT"
+    with * have "s \<le> Y + (-X) + 1" by simp
+    also have "\<dots> < Y + (-X) + 2" by simp
+    finally show "s < Y + (-X) + 2" .
+  qed
   moreover have "\<forall>y \<in> ?SHIFT. 0 < y" by auto
   moreover have shifted_not_empty: "\<exists>u. u \<in> ?SHIFT"
     using X_in_S and Y_isUb by auto
-  ultimately obtain t where t_is_Lub: "isLub UNIV ?SHIFT t"
-    using posreals_complete [of ?SHIFT] by blast
+  ultimately obtain t where t_is_Lub: "\<forall>y. (\<exists>x\<in>?SHIFT. y < x) = (y < t)"
+    using posreal_complete [of ?SHIFT] unfolding bdd_above_def by blast
 
   show ?thesis
   proof
-    show "isLub UNIV S (t + X + (-1))"
-    proof (rule isLubI2)
-      {
-        fix x
-        assume "isUb (UNIV::real set) S x"
-        hence "isUb (UNIV::real set) (?SHIFT) (x + (-X) + 1)"
-          using S_Ub_is_SHIFT_Ub by simp
-        hence "t \<le> (x + (-X) + 1)"
-          using t_is_Lub by (simp add: isLub_le_isUb)
-        hence "t + X + -1 \<le> x" by arith
-      }
-      then show "(t + X + -1) <=* Collect (isUb UNIV S)"
-        by (simp add: setgeI)
+    show "(\<forall>s\<in>S. s \<le> (t + X + (-1))) \<and> (\<forall>y. (\<forall>s\<in>S. s \<le> y) \<longrightarrow> (t + X + (-1)) \<le> y)"
+    proof safe
+      fix x
+      assume "\<forall>s\<in>S. s \<le> x"
+      hence "\<forall>s\<in>?SHIFT. s \<le> x + (-X) + 1"
+        using S_Ub_is_SHIFT_Ub by simp
+      then have "\<not> x + (-X) + 1 < t"
+        by (subst t_is_Lub[rule_format, symmetric]) (simp add: not_less)
+      thus "t + X + -1 \<le> x" by arith
     next
-      show "isUb UNIV S (t + X + -1)"
-      proof -
-        {
-          fix y
-          assume y_in_S: "y \<in> S"
-          have "y \<le> t + X + -1"
-          proof -
-            obtain "u" where u_in_shift: "u \<in> ?SHIFT" using shifted_not_empty ..
-            hence "\<exists> x \<in> S. u = x + - X + 1" by simp
-            then obtain "x" where x_and_u: "u = x + - X + 1" ..
-            have u_le_t: "u \<le> t" using u_in_shift and t_is_Lub by (simp add: isLubD2)
+      fix y
+      assume y_in_S: "y \<in> S"
+      obtain "u" where u_in_shift: "u \<in> ?SHIFT" using shifted_not_empty ..
+      hence "\<exists> x \<in> S. u = x + - X + 1" by simp
+      then obtain "x" where x_and_u: "u = x + - X + 1" ..
+      have u_le_t: "u \<le> t"
+      proof (rule dense_le)
+        fix x assume "x < u" then have "x < t"
+          using u_in_shift t_is_Lub by auto
+        then show "x \<le> t"  by simp
+      qed
 
-            show ?thesis
-            proof cases
-              assume "y \<le> x"
-              moreover have "x = u + X + - 1" using x_and_u by arith
-              moreover have "u + X + - 1  \<le> t + X + -1" using u_le_t by arith
-              ultimately show "y  \<le> t + X + -1" by arith
-            next
-              assume "~(y \<le> x)"
-              hence x_less_y: "x < y" by arith
+      show "y \<le> t + X + -1"
+      proof cases
+        assume "y \<le> x"
+        moreover have "x = u + X + - 1" using x_and_u by arith
+        moreover have "u + X + - 1  \<le> t + X + -1" using u_le_t by arith
+        ultimately show "y  \<le> t + X + -1" by arith
+      next
+        assume "~(y \<le> x)"
+        hence x_less_y: "x < y" by arith
 
-              have "x + (-X) + 1 \<in> ?SHIFT" using x_and_u and u_in_shift by simp
-              hence "0 < x + (-X) + 1" by simp
-              hence "0 < y + (-X) + 1" using x_less_y by arith
-              hence "y + (-X) + 1 \<in> ?SHIFT" using y_in_S by simp
-              hence "y + (-X) + 1 \<le> t" using t_is_Lub  by (simp add: isLubD2)
-              thus ?thesis by simp
-            qed
-          qed
-        }
-        then show ?thesis by (simp add: isUb_def setle_def)
+        have "x + (-X) + 1 \<in> ?SHIFT" using x_and_u and u_in_shift by simp
+        hence "0 < x + (-X) + 1" by simp
+        hence "0 < y + (-X) + 1" using x_less_y by arith
+        hence *: "y + (-X) + 1 \<in> ?SHIFT" using y_in_S by simp
+        have "y + (-X) + 1 \<le> t"
+        proof (rule dense_le)
+          fix x assume "x < y + (-X) + 1" then have "x < t"
+            using * t_is_Lub by auto
+          then show "x \<le> t"  by simp
+        qed
+        thus ?thesis by simp
       qed
     qed
   qed
 qed
 
-text{*A version of the same theorem without all those predicates!*}
-lemma reals_complete2:
-  fixes S :: "(real set)"
-  assumes "\<exists>y. y\<in>S" and "\<exists>(x::real). \<forall>y\<in>S. y \<le> x"
-  shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) & 
-               (\<forall>z. ((\<forall>y\<in>S. y \<le> z) --> x \<le> z))"
-proof -
-  have "\<exists>x. isLub UNIV S x" 
-    by (rule reals_complete)
-       (auto simp add: isLub_def isUb_def leastP_def setle_def setge_def assms)
-  thus ?thesis
-    by (metis UNIV_I isLub_isUb isLub_le_isUb isUbD isUb_def setleI)
-qed
-
-
 subsection {* The Archimedean Property of the Reals *}
 
 theorem reals_Archimedean:
@@ -1969,34 +1871,30 @@
       by (rule mult_right_mono)
     thus "x * of_nat (Suc n) \<le> 1" by (simp del: of_nat_Suc)
   qed
-  hence "{z. \<exists>n. z = x * (of_nat (Suc n))} *<= 1"
-    by (simp add: setle_def del: of_nat_Suc, safe, rule spec)
-  hence "isUb (UNIV::real set) {z. \<exists>n. z = x * (of_nat (Suc n))} 1"
-    by (simp add: isUbI)
-  hence "\<exists>Y. isUb (UNIV::real set) {z. \<exists>n. z = x* (of_nat (Suc n))} Y" ..
-  moreover have "\<exists>X. X \<in> {z. \<exists>n. z = x* (of_nat (Suc n))}" by auto
-  ultimately have "\<exists>t. isLub UNIV {z. \<exists>n. z = x * of_nat (Suc n)} t"
-    by (simp add: reals_complete)
-  then obtain "t" where
-    t_is_Lub: "isLub UNIV {z. \<exists>n. z = x * of_nat (Suc n)} t" ..
+  hence 2: "bdd_above {z. \<exists>n. z = x * (of_nat (Suc n))}"
+    by (auto intro!: bdd_aboveI[of _ 1])
+  have 1: "\<exists>X. X \<in> {z. \<exists>n. z = x* (of_nat (Suc n))}" by auto
+  obtain t where
+    upper: "\<And>z. z \<in> {z. \<exists>n. z = x * of_nat (Suc n)} \<Longrightarrow> z \<le> t" and
+    least: "\<And>y. (\<And>a. a \<in> {z. \<exists>n. z = x * of_nat (Suc n)} \<Longrightarrow> a \<le> y) \<Longrightarrow> t \<le> y"
+    using reals_complete[OF 1 2] by auto
+
 
-  have "\<forall>n::nat. x * of_nat n \<le> t + - x"
-  proof
-    fix n
-    from t_is_Lub have "x * of_nat (Suc n) \<le> t"
-      by (simp add: isLubD2)
-    hence  "x * (of_nat n) + x \<le> t"
-      by (simp add: distrib_left)
-    thus  "x * (of_nat n) \<le> t + - x" by arith
+  have "t \<le> t + - x"
+  proof (rule least)
+    fix a assume a: "a \<in> {z. \<exists>n. z = x * (of_nat (Suc n))}"
+    have "\<forall>n::nat. x * of_nat n \<le> t + - x"
+    proof
+      fix n
+      have "x * of_nat (Suc n) \<le> t"
+        by (simp add: upper)
+      hence  "x * (of_nat n) + x \<le> t"
+        by (simp add: distrib_left)
+      thus  "x * (of_nat n) \<le> t + - x" by arith
+    qed    hence "\<forall>m. x * of_nat (Suc m) \<le> t + - x" by (simp del: of_nat_Suc)
+    with a show "a \<le> t + - x"
+      by auto
   qed
-
-  hence "\<forall>m. x * of_nat (Suc m) \<le> t + - x" by (simp del: of_nat_Suc)
-  hence "{z. \<exists>n. z = x * (of_nat (Suc n))}  *<= (t + - x)"
-    by (auto simp add: setle_def)
-  hence "isUb (UNIV::real set) {z. \<exists>n. z = x * (of_nat (Suc n))} (t + (-x))"
-    by (simp add: isUbI)
-  hence "t \<le> t + - x"
-    using t_is_Lub by (simp add: isLub_le_isUb)
   thus False using x_pos by arith
 qed
 
--- a/src/HOL/ex/Gauge_Integration.thy	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/HOL/ex/Gauge_Integration.thy	Mon Nov 11 17:44:21 2013 +0100
@@ -511,9 +511,9 @@
   case False
   then have "inverse (z - x) * (f z - f x - f' x * (z - x)) = (f z - f x) / (z - x) - f' x"
     apply (subst mult_commute)
-    apply (simp add: distrib_right diff_minus)
+    apply (simp add: left_diff_distrib)
     apply (simp add: mult_assoc divide_inverse)
-    apply (simp add: distrib_right)
+    apply (simp add: ring_distribs)
     done
   moreover from False `\<bar>z - x\<bar> < s` have "\<bar>(f z - f x) / (z - x) - f' x\<bar> < e / 2"
     by (rule P)
--- a/src/Provers/splitter.ML	Mon Nov 11 17:34:44 2013 +0100
+++ b/src/Provers/splitter.ML	Mon Nov 11 17:44:21 2013 +0100
@@ -79,6 +79,8 @@
   fold add_thm splits []
 end;
 
+val abss = fold (Term.abs o pair "");
+
 (* ------------------------------------------------------------------------- *)
 (* mk_case_split_tac                                                         *)
 (* ------------------------------------------------------------------------- *)
@@ -100,31 +102,36 @@
   (Syntax.read_prop_global Pure.thy "P(%x. Q(x)) == P(%x. R(x))")
   (fn {prems, ...} => rewrite_goals_tac prems THEN rtac reflexive_thm 1)
 
+val _ $ _ $ (_ $ (_ $ abs_lift) $ _) = prop_of lift;
+
 val trlift = lift RS transitive_thm;
-val _ $ (P $ _) $ _ = concl_of trlift;
 
 
 (************************************************************************
    Set up term for instantiation of P in the lift-theorem
 
-   Ts    : types of parameters (i.e. variables bound by meta-quantifiers)
    t     : lefthand side of meta-equality in subgoal
            the lift theorem is applied to (see select)
    pos   : "path" leading to abstraction, coded as a list
    T     : type of body of P(...)
-   maxi  : maximum index of Vars
 *************************************************************************)
 
-fun mk_cntxt Ts t pos T maxi =
-  let fun var (t,i) = Var(("X",i),type_of1(Ts,t));
-      fun down [] t i = Bound 0
-        | down (p::ps) t i =
-            let val (h,ts) = strip_comb t
-                val v1 = ListPair.map var (take p ts, i upto (i+p-1))
-                val u::us = drop p ts
-                val v2 = ListPair.map var (us, (i+p) upto (i+length(ts)-2))
-      in list_comb(h,v1@[down ps u (i+length ts)]@v2) end;
-  in Abs("", T, down (rev pos) t maxi) end;
+fun mk_cntxt t pos T =
+  let
+    fun down [] t = (Bound 0, t)
+      | down (p :: ps) t =
+          let
+            val (h, ts) = strip_comb t
+            val (ts1, u :: ts2) = chop p ts
+            val (u1, u2) = down ps u
+          in
+            (list_comb (incr_boundvars 1 h,
+               map (incr_boundvars 1) ts1 @ u1 ::
+               map (incr_boundvars 1) ts2),
+             u2)
+          end;
+    val (u1, u2) = down (rev pos) t
+  in (Abs ("", T, u1), u2) end;
 
 
 (************************************************************************
@@ -301,15 +308,18 @@
              the split theorem is applied to (see cmap)
    T,U,pos : see mk_split_pack
    state   : current proof state
-   lift    : the lift theorem
    i       : no. of subgoal
 **************************************************************)
 
 fun inst_lift Ts t (T, U, pos) state i =
   let
     val cert = cterm_of (Thm.theory_of_thm state);
-    val cntxt = mk_cntxt Ts t pos (T --> U) (Thm.maxidx_of trlift);
-  in cterm_instantiate [(cert P, cert cntxt)] trlift
+    val (cntxt, u) = mk_cntxt t pos (T --> U);
+    val trlift' = Thm.lift_rule (Thm.cprem_of state i)
+      (Thm.rename_boundvars abs_lift u trlift);
+    val (P, _) = strip_comb (fst (Logic.dest_equals
+      (Logic.strip_assums_concl (Thm.prop_of trlift'))));
+  in cterm_instantiate [(cert P, cert (abss Ts cntxt))] trlift'
   end;
 
 
@@ -333,7 +343,6 @@
       (Logic.strip_assums_concl (Thm.prop_of thm'))));
     val cert = cterm_of (Thm.theory_of_thm state);
     val cntxt = mk_cntxt_splitthm t tt TB;
-    val abss = fold (fn T => fn t => Abs ("", T, t));
   in cterm_instantiate [(cert P, cert (abss Ts cntxt))] thm'
   end;
 
@@ -348,7 +357,7 @@
 fun split_tac [] i = no_tac
   | split_tac splits i =
   let val cmap = cmap_of_split_thms splits
-      fun lift_tac Ts t p st = rtac (inst_lift Ts t p st i) i st
+      fun lift_tac Ts t p st = compose_tac (false, inst_lift Ts t p st i, 2) i st
       fun lift_split_tac state =
             let val (Ts, t, splits) = select cmap state i
             in case splits of