merged
authorhuffman
Mon, 22 Feb 2010 11:19:15 -0800
changeset 35289 08e11c587c3d
parent 35288 aa7da51ae1ef (current diff)
parent 35285 40da65ef68e3 (diff)
child 35290 3707f625314f
merged
src/HOLCF/IOA/meta_theory/Seq.thy
--- a/Admin/isatest/isatest-makedist	Mon Feb 22 11:17:41 2010 -0800
+++ b/Admin/isatest/isatest-makedist	Mon Feb 22 11:19:15 2010 -0800
@@ -94,12 +94,6 @@
 $SSH macbroy22 "$MAKEALL $HOME/settings/at-poly"
 # give test some time to copy settings and start
 sleep 15
-$SSH macbroy20 "$MAKEALL $HOME/settings/at-poly-5.1-para-e"
-sleep 15
-#$SSH macbroy24 "$MAKEALL -l HOL proofterms $HOME/settings/at-sml-dev-p"
-#sleep 15
-$SSH macbroy21 "$MAKEALL $HOME/settings/at64-poly-5.1-para"
-sleep 15
 $SSH macbroy23 "$MAKEALL -l HOL HOL-ex $HOME/settings/at-sml-dev-e"
 sleep 15
 $SSH macbroy24 "$MAKEALL $HOME/settings/at64-poly"
--- a/Admin/isatest/settings/at64-poly	Mon Feb 22 11:17:41 2010 -0800
+++ b/Admin/isatest/settings/at64-poly	Mon Feb 22 11:19:15 2010 -0800
@@ -1,12 +1,12 @@
 # -*- shell-script -*- :mode=shellscript:
 
-  POLYML_HOME="/home/polyml/polyml-5.2"
-  ML_SYSTEM="polyml-5.2"
+  POLYML_HOME="/home/polyml/polyml-5.3.0"
+  ML_SYSTEM="polyml-5.3.0"
   ML_PLATFORM="x86_64-linux"
   ML_HOME="$POLYML_HOME/$ML_PLATFORM"
-  ML_OPTIONS="-H 500"
+  ML_OPTIONS="-H 1000"
 
-ISABELLE_HOME_USER=~/isabelle-at64-poly-e
+ISABELLE_HOME_USER=~/isabelle-at64-poly
 
 # Where to look for isabelle tools (multiple dirs separated by ':').
 ISABELLE_TOOLS="$ISABELLE_HOME/lib/Tools"
@@ -24,5 +24,4 @@
 
 ISABELLE_USEDIR_OPTIONS="-i true -d pdf -v true"
 
-unset KODKODI
-
+init_component /home/isabelle/contrib_devel/kodkodi
--- a/NEWS	Mon Feb 22 11:17:41 2010 -0800
+++ b/NEWS	Mon Feb 22 11:19:15 2010 -0800
@@ -4,6 +4,41 @@
 New in this Isabelle version
 ----------------------------
 
+*** General ***
+
+* Authentic syntax for *all* term constants: provides simple and
+robust correspondence between formal entities and concrete syntax.
+Substantial INCOMPATIBILITY concerning low-level syntax translations
+(translation rules and translation functions in ML).  Some hints on
+upgrading:
+
+  - Many existing uses of 'syntax' and 'translations' can be replaced
+    by more modern 'notation' and 'abbreviation', which are
+    independent of this issue.
+
+  - 'translations' require markup within the AST; the term syntax
+    provides the following special forms:
+
+      CONST c   -- produces syntax version of constant c from context
+      XCONST c  -- literally c, checked as constant from context
+      c         -- literally c, if declared by 'syntax'
+
+    Plain identifiers are treated as AST variables -- occasionally the
+    system indicates accidental variables via the error "rhs contains
+    extra variables".
+
+  - 'parse_translation' etc. in ML may use the following
+    antiquotations:
+
+      @{const_syntax c}   -- ML version of "CONST c" above
+      @{syntax_const c}   -- literally c (checked wrt. 'syntax' declarations)
+
+Note that old non-authentic syntax was based on unqualified base
+names, so all of the above would coincide.  Recall that 'print_syntax'
+and ML_command "set Syntax.trace_ast" help to diagnose syntax
+problems.
+
+
 *** Pure ***
 
 * Code generator: details of internal data cache have no impact on
@@ -28,7 +63,7 @@
 * Some generic constants have been put to appropriate theories:
 
     less_eq, less: Orderings
-    abs, sgn: Groups
+    zero, one, plus, minus, uminus, times, abs, sgn: Groups
     inverse, divide: Rings
 
 INCOMPATIBILITY.
@@ -88,8 +123,7 @@
 INCOMPATIBILITY.
 
 * New theory Algebras contains generic algebraic structures and
-generic algebraic operations.  INCOMPATIBILTY: constants zero, one,
-plus, minus, uminus and times have been moved from HOL.thy to Algebras.thy.
+generic algebraic operations.
 
 * HOLogic.strip_psplit: types are returned in syntactic order, similar
 to other strip and tuple operations.  INCOMPATIBILITY.
@@ -98,10 +132,14 @@
 replaced by new-style primrec, especially in theory List.  The corresponding
 constants now have authentic syntax.  INCOMPATIBILITY.
 
-* Reorganized theory Multiset: less duplication, less historical
-organization of sections, conversion from associations lists to
-multisets, rudimentary code generation.  Use insert_DiffM2 [symmetric]
-instead of elem_imp_eq_diff_union, if needed.  INCOMPATIBILITY.
+* Reorganized theory Multiset: swapped notation of pointwise and multiset order:
+  * pointwise ordering is instance of class order with standard syntax <= and <;
+  * multiset ordering has syntax <=# and <#; partial order properties are provided
+      by means of interpretation with prefix multiset_order.
+Less duplication, less historical organization of sections,
+conversion from associations lists to multisets, rudimentary code generation.
+Use insert_DiffM2 [symmetric] instead of elem_imp_eq_diff_union, if needed.
+INCOMPATIBILITY.
 
 * Reorganized theory Sum_Type; Inl and Inr now have authentic syntax.
 INCOMPATIBILITY.
--- a/doc-src/Classes/Thy/Classes.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/doc-src/Classes/Thy/Classes.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -362,17 +362,18 @@
 
 text {*
   \noindent The connection to the type system is done by means
-  of a primitive axclass
+  of a primitive type class
 *} (*<*)setup %invisible {* Sign.add_path "foo" *}
 (*>*)
-axclass %quote idem < type
-  idem: "f (f x) = f x" (*<*)setup %invisible {* Sign.parent_path *}(*>*)
+classes %quote idem < type
+(*<*)axiomatization where idem: "f (f (x::\<alpha>\<Colon>idem)) = f x"
+setup %invisible {* Sign.parent_path *}(*>*)
 
 text {* \noindent together with a corresponding interpretation: *}
 
 interpretation %quote idem_class:
   idem "f \<Colon> (\<alpha>\<Colon>idem) \<Rightarrow> \<alpha>"
-proof qed (rule idem)
+(*<*)proof qed (rule idem)(*>*)
 
 text {*
   \noindent This gives you the full power of the Isabelle module system;
@@ -414,8 +415,7 @@
 subsection {* Derived definitions *}
 
 text {*
-  Isabelle locales support a concept of local definitions
-  in locales:
+  Isabelle locales are targets which support local definitions:
 *}
 
 primrec %quote (in monoid) pow_nat :: "nat \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>" where
--- a/doc-src/Classes/Thy/document/Classes.tex	Mon Feb 22 11:17:41 2010 -0800
+++ b/doc-src/Classes/Thy/document/Classes.tex	Mon Feb 22 11:19:15 2010 -0800
@@ -641,7 +641,7 @@
 %
 \begin{isamarkuptext}%
 \noindent The connection to the type system is done by means
-  of a primitive axclass%
+  of a primitive type class%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \ %
@@ -663,9 +663,8 @@
 \endisadelimquote
 %
 \isatagquote
-\isacommand{axclass}\isamarkupfalse%
-\ idem\ {\isacharless}\ type\isanewline
-\ \ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}\ %
+\isacommand{classes}\isamarkupfalse%
+\ idem\ {\isacharless}\ type%
 \endisatagquote
 {\isafoldquote}%
 %
@@ -698,10 +697,7 @@
 \isatagquote
 \isacommand{interpretation}\isamarkupfalse%
 \ idem{\isacharunderscore}class{\isacharcolon}\isanewline
-\ \ idem\ {\isachardoublequoteopen}f\ {\isasymColon}\ {\isacharparenleft}{\isasymalpha}{\isasymColon}idem{\isacharparenright}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline
-\isacommand{proof}\isamarkupfalse%
-\ \isacommand{qed}\isamarkupfalse%
-\ {\isacharparenleft}rule\ idem{\isacharparenright}%
+\ \ idem\ {\isachardoublequoteopen}f\ {\isasymColon}\ {\isacharparenleft}{\isasymalpha}{\isasymColon}idem{\isacharparenright}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}%
 \endisatagquote
 {\isafoldquote}%
 %
--- a/doc-src/IsarRef/Thy/Spec.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/doc-src/IsarRef/Thy/Spec.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -738,44 +738,6 @@
 *}
 
 
-subsection {* Old-style axiomatic type classes \label{sec:axclass} *}
-
-text {*
-  \begin{matharray}{rcl}
-    @{command_def "axclass"} & : & @{text "theory \<rightarrow> theory"}
-  \end{matharray}
-
-  Axiomatic type classes are Isabelle/Pure's primitive
-  interface to type classes.  For practical
-  applications, you should consider using classes
-  (cf.~\secref{sec:classes}) which provide high level interface.
-
-  \begin{rail}
-    'axclass' classdecl (axmdecl prop +)
-    ;
-  \end{rail}
-
-  \begin{description}
-  
-  \item @{command "axclass"}~@{text "c \<subseteq> c\<^sub>1, \<dots>, c\<^sub>n axms"} defines an
-  axiomatic type class as the intersection of existing classes, with
-  additional axioms holding.  Class axioms may not contain more than
-  one type variable.  The class axioms (with implicit sort constraints
-  added) are bound to the given names.  Furthermore a class
-  introduction rule is generated (being bound as @{text
-  c_class.intro}); this rule is employed by method @{method
-  intro_classes} to support instantiation proofs of this class.
-  
-  The ``class axioms'' (which are derived from the internal class
-  definition) are stored as theorems according to the given name
-  specifications; the name space prefix @{text "c_class"} is added
-  here.  The full collection of these facts is also stored as @{text
-  c_class.axioms}.
-  
-  \end{description}
-*}
-
-
 section {* Unrestricted overloading *}
 
 text {*
@@ -962,8 +924,9 @@
 
   \item @{command "classrel"}~@{text "c\<^sub>1 \<subseteq> c\<^sub>2"} states subclass
   relations between existing classes @{text "c\<^sub>1"} and @{text "c\<^sub>2"}.
-  This is done axiomatically!  The @{command_ref "instance"} command
-  (see \secref{sec:axclass}) provides a way to introduce proven class
+  This is done axiomatically!  The @{command_ref "instance"}
+  and @{command_ref "subclass"} command
+  (see \secref{sec:class}) provide a way to introduce proven class
   relations.
 
   \item @{command "defaultsort"}~@{text s} makes sort @{text s} the
@@ -1021,8 +984,8 @@
 
   \item @{command "arities"}~@{text "t :: (s\<^sub>1, \<dots>, s\<^sub>n)s"} augments
   Isabelle's order-sorted signature of types by new type constructor
-  arities.  This is done axiomatically!  The @{command_ref "instance"}
-  command (see \secref{sec:axclass}) provides a way to introduce
+  arities.  This is done axiomatically!  The @{command_ref "instantiation"}
+  target (see \secref{sec:class}) provides a way to introduce
   proven type arities.
 
   \end{description}
--- a/doc-src/IsarRef/Thy/document/Spec.tex	Mon Feb 22 11:17:41 2010 -0800
+++ b/doc-src/IsarRef/Thy/document/Spec.tex	Mon Feb 22 11:19:15 2010 -0800
@@ -755,43 +755,6 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
-\isamarkupsubsection{Old-style axiomatic type classes \label{sec:axclass}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-\begin{matharray}{rcl}
-    \indexdef{}{command}{axclass}\hypertarget{command.axclass}{\hyperlink{command.axclass}{\mbox{\isa{\isacommand{axclass}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}}
-  \end{matharray}
-
-  Axiomatic type classes are Isabelle/Pure's primitive
-  interface to type classes.  For practical
-  applications, you should consider using classes
-  (cf.~\secref{sec:classes}) which provide high level interface.
-
-  \begin{rail}
-    'axclass' classdecl (axmdecl prop +)
-    ;
-  \end{rail}
-
-  \begin{description}
-  
-  \item \hyperlink{command.axclass}{\mbox{\isa{\isacommand{axclass}}}}~\isa{{\isachardoublequote}c\ {\isasymsubseteq}\ c\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ c\isactrlsub n\ axms{\isachardoublequote}} defines an
-  axiomatic type class as the intersection of existing classes, with
-  additional axioms holding.  Class axioms may not contain more than
-  one type variable.  The class axioms (with implicit sort constraints
-  added) are bound to the given names.  Furthermore a class
-  introduction rule is generated (being bound as \isa{c{\isacharunderscore}class{\isachardot}intro}); this rule is employed by method \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isacharunderscore}classes}}} to support instantiation proofs of this class.
-  
-  The ``class axioms'' (which are derived from the internal class
-  definition) are stored as theorems according to the given name
-  specifications; the name space prefix \isa{{\isachardoublequote}c{\isacharunderscore}class{\isachardoublequote}} is added
-  here.  The full collection of these facts is also stored as \isa{c{\isacharunderscore}class{\isachardot}axioms}.
-  
-  \end{description}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
 \isamarkupsection{Unrestricted overloading%
 }
 \isamarkuptrue%
@@ -996,8 +959,9 @@
 
   \item \hyperlink{command.classrel}{\mbox{\isa{\isacommand{classrel}}}}~\isa{{\isachardoublequote}c\isactrlsub {\isadigit{1}}\ {\isasymsubseteq}\ c\isactrlsub {\isadigit{2}}{\isachardoublequote}} states subclass
   relations between existing classes \isa{{\isachardoublequote}c\isactrlsub {\isadigit{1}}{\isachardoublequote}} and \isa{{\isachardoublequote}c\isactrlsub {\isadigit{2}}{\isachardoublequote}}.
-  This is done axiomatically!  The \indexref{}{command}{instance}\hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} command
-  (see \secref{sec:axclass}) provides a way to introduce proven class
+  This is done axiomatically!  The \indexref{}{command}{instance}\hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}}
+  and \indexref{}{command}{subclass}\hyperlink{command.subclass}{\mbox{\isa{\isacommand{subclass}}}} command
+  (see \secref{sec:class}) provide a way to introduce proven class
   relations.
 
   \item \hyperlink{command.defaultsort}{\mbox{\isa{\isacommand{defaultsort}}}}~\isa{s} makes sort \isa{s} the
@@ -1056,8 +1020,8 @@
 
   \item \hyperlink{command.arities}{\mbox{\isa{\isacommand{arities}}}}~\isa{{\isachardoublequote}t\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}s\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ s\isactrlsub n{\isacharparenright}s{\isachardoublequote}} augments
   Isabelle's order-sorted signature of types by new type constructor
-  arities.  This is done axiomatically!  The \indexref{}{command}{instance}\hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}}
-  command (see \secref{sec:axclass}) provides a way to introduce
+  arities.  This is done axiomatically!  The \indexref{}{command}{instantiation}\hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}}
+  target (see \secref{sec:class}) provides a way to introduce
   proven type arities.
 
   \end{description}%
--- a/doc-src/Main/Docs/Main_Doc.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/doc-src/Main/Docs/Main_Doc.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -48,8 +48,8 @@
 \smallskip
 
 \begin{supertabular}{@ {} l @ {~::~} l l @ {}}
-@{const Algebras.less_eq} & @{typeof Algebras.less_eq} & (\verb$<=$)\\
-@{const Algebras.less} & @{typeof Algebras.less}\\
+@{const Orderings.less_eq} & @{typeof Orderings.less_eq} & (\verb$<=$)\\
+@{const Orderings.less} & @{typeof Orderings.less}\\
 @{const Orderings.Least} & @{typeof Orderings.Least}\\
 @{const Orderings.min} & @{typeof Orderings.min}\\
 @{const Orderings.max} & @{typeof Orderings.max}\\
--- a/doc-src/Main/Docs/document/Main_Doc.tex	Mon Feb 22 11:17:41 2010 -0800
+++ b/doc-src/Main/Docs/document/Main_Doc.tex	Mon Feb 22 11:19:15 2010 -0800
@@ -416,7 +416,7 @@
 \begin{supertabular}{@ {} l @ {\quad$\equiv$\quad} l l @ {}}
 \isa{{\isasymSum}A} & \isa{{\isachardoublequote}setsum\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ x{\isacharparenright}\ A{\isachardoublequote}} & (\verb$SUM$)\\
 \isa{{\isasymSum}x{\isasymin}A{\isachardot}\ t} & \isa{{\isachardoublequote}setsum\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ t{\isacharparenright}\ A{\isachardoublequote}}\\
-\isa{{\isachardoublequote}{\isasymSum}x{\isacharbar}P{\isachardot}\ t{\isachardoublequote}} & \isa{{\isasymSum}x{\isasymin}{\isacharbraceleft}x{\isachardot}\ P{\isacharbraceright}{\isachardot}\ t}\\
+\isa{{\isachardoublequote}{\isasymSum}x{\isacharbar}P{\isachardot}\ t{\isachardoublequote}} & \isa{{\isasymSum}x\ {\isacharbar}\ P{\isachardot}\ t}\\
 \multicolumn{2}{@ {}l@ {}}{Similarly for \isa{{\isasymProd}} instead of \isa{{\isasymSum}}} & (\verb$PROD$)\\
 \end{supertabular}
 
--- a/doc-src/Nitpick/nitpick.tex	Mon Feb 22 11:17:41 2010 -0800
+++ b/doc-src/Nitpick/nitpick.tex	Mon Feb 22 11:19:15 2010 -0800
@@ -141,11 +141,11 @@
 
 This section introduces Nitpick by presenting small examples. If possible, you
 should try out the examples on your workstation. Your theory file should start
-the standard way:
+as follows:
 
 \prew
 \textbf{theory}~\textit{Scratch} \\
-\textbf{imports}~\textit{Main} \\
+\textbf{imports}~\textit{Main~Coinductive\_List~Quotient\_Product~RealDef} \\
 \textbf{begin}
 \postw
 
@@ -439,7 +439,7 @@
 value (displayed as `$\unk$'). The type \textit{int} is handled similarly.
 Internally, undefined values lead to a three-valued logic.
 
-Here is an example involving \textit{int}:
+Here is an example involving \textit{int\/}:
 
 \prew
 \textbf{lemma} ``$\lbrakk i \le j;\> n \le (m{\Colon}\mathit{int})\rbrakk \,\Longrightarrow\, i * n + j * m \le i * m + j * n$'' \\
@@ -627,7 +627,7 @@
 genuine counterexample for \textit{card}~$'a~\textit{list}$~= 10, it is very
 unlikely that one could be found for smaller cardinalities.
 
-\subsection{Typedefs, Records, Rationals, and Reals}
+\subsection{Typedefs, Quotient Types, Records, Rationals, and Reals}
 \label{typedefs-records-rationals-and-reals}
 
 Nitpick generally treats types declared using \textbf{typedef} as datatypes
@@ -651,12 +651,41 @@
 \hbox{}\qquad\qquad $\textit{three} = \{\Abs{0},\, \Abs{1},\, \Abs{2},\, \unr\}$
 \postw
 
-%% MARK
 In the output above, $\Abs{n}$ abbreviates $\textit{Abs\_three}~n$.
 
-%% MARK
-Records, which are implemented as \textbf{typedef}s behind the scenes, are
-handled in much the same way:
+Quotient types are handled in much the same way. The following fragment defines
+the integer type \textit{my\_int} by encoding the integer $x$ by a pair of
+natural numbers $(m, n)$ such that $x + n = m$:
+
+\prew
+\textbf{fun} \textit{my\_int\_rel} \textbf{where} \\
+``$\textit{my\_int\_rel}~(x,\, y)~(u,\, v) = (x + v = u + y)$'' \\[2\smallskipamount]
+%
+\textbf{quotient\_type}~\textit{my\_int} = ``$\textit{nat} \times \textit{nat\/}$''$\;{/}\;$\textit{my\_int\_rel} \\
+\textbf{by}~(\textit{auto simp add\/}:\ \textit{equivp\_def expand\_fun\_eq}) \\[2\smallskipamount]
+%
+\textbf{definition}~\textit{add\_raw}~\textbf{where} \\
+``$\textit{add\_raw} \,\equiv\, \lambda(x,\, y)~(u,\, v).\; (x + (u\Colon\textit{nat}), y + (v\Colon\textit{nat}))$'' \\[2\smallskipamount]
+%
+\textbf{quotient\_definition} ``$\textit{add\/}\Colon\textit{my\_int} \Rightarrow \textit{my\_int} \Rightarrow \textit{my\_int\/}$'' \textbf{is} \textit{add\_raw} \\[2\smallskipamount]
+%
+\textbf{lemma} ``$\textit{add}~x~y = \textit{add}~x~x$'' \\
+\textbf{nitpick} [\textit{show\_datatypes}] \\[2\smallskipamount]
+\slshape Nitpick found a counterexample: \\[2\smallskipamount]
+\hbox{}\qquad Free variables: \nopagebreak \\
+\hbox{}\qquad\qquad $x = \Abs{(0,\, 0)}$ \\
+\hbox{}\qquad\qquad $y = \Abs{(1,\, 0)}$ \\
+\hbox{}\qquad Datatypes: \\
+\hbox{}\qquad\qquad $\textit{nat} = \{0,\, 1,\, \unr\}$ \\
+\hbox{}\qquad\qquad $\textit{nat} \times \textit{nat} = \{(0,\, 0),\> (1,\, 0),\> \unr\}$ \\
+\hbox{}\qquad\qquad $\textit{my\_int} = \{\Abs{(0,\, 0)},\> \Abs{(1,\, 0)},\> \unr\}$
+\postw
+
+In the counterexample, $\Abs{(0,\, 0)}$ and $\Abs{(1,\, 0)}$ represent the
+integers $0$ and $1$, respectively. Other representants would have been
+possible---e.g., $\Abs{(5,\, 5)}$ and $\Abs{(12,\, 11)}$.
+
+Records are also handled as datatypes with a single constructor:
 
 \prew
 \textbf{record} \textit{point} = \\
@@ -675,6 +704,8 @@
 & \lparr\textit{Xcoord} = 1,\> \textit{Ycoord} = 1\rparr,\, \unr\}\end{aligned}$
 \postw
 
+
+
 Finally, Nitpick provides rudimentary support for rationals and reals using a
 similar approach:
 
@@ -949,13 +980,13 @@
 \slshape The type ``\kern1pt$'a$'' passed the monotonicity test. Nitpick might be able to skip
 some scopes. \\[2\smallskipamount]
 Trying 8 scopes: \\
-\hbox{}\qquad \textit{card} $'a$~= 1, \textit{card} ``\kern1pt$'a~\textit{list}$''~= 1,
+\hbox{}\qquad \textit{card} $'a$~= 1, \textit{card} ``\kern1pt$'a~\textit{list\/}$''~= 1,
 and \textit{bisim\_depth}~= 0. \\
 \hbox{}\qquad $\qquad\vdots$ \\[.5\smallskipamount]
-\hbox{}\qquad \textit{card} $'a$~= 8, \textit{card} ``\kern1pt$'a~\textit{list}$''~= 8,
+\hbox{}\qquad \textit{card} $'a$~= 8, \textit{card} ``\kern1pt$'a~\textit{list\/}$''~= 8,
 and \textit{bisim\_depth}~= 7. \\[2\smallskipamount]
 Nitpick found a counterexample for {\itshape card}~$'a$ = 2,
-\textit{card}~``\kern1pt$'a~\textit{list}$''~= 2, and \textit{bisim\_\allowbreak
+\textit{card}~``\kern1pt$'a~\textit{list\/}$''~= 2, and \textit{bisim\_\allowbreak
 depth}~= 1:
 \\[2\smallskipamount]
 \hbox{}\qquad Free variables: \nopagebreak \\
@@ -1118,7 +1149,7 @@
 \textit{Lam}~(\textit{Lam}~(\textit{Var}~0))$. Using the traditional
 $\lambda$-term notation, $t$~is
 $\lambda x\, y.\> x$ whereas $\textit{subst}~\sigma~t$ is $\lambda x\, y.\> y$.
-The bug is in \textit{subst}: The $\textit{lift}~(\sigma~m)~1$ call should be
+The bug is in \textit{subst\/}: The $\textit{lift}~(\sigma~m)~1$ call should be
 replaced with $\textit{lift}~(\sigma~m)~0$.
 
 An interesting aspect of Nitpick's verbose output is that it assigned inceasing
@@ -1509,7 +1540,7 @@
 completeness of the set $S$. First, soundness:
 
 \prew
-\textbf{theorem}~\textit{S\_sound}: \\
+\textbf{theorem}~\textit{S\_sound\/}: \\
 ``$w \in S \longrightarrow \textit{length}~[x\mathbin{\leftarrow} w.\; x = a] =
   \textit{length}~[x\mathbin{\leftarrow} w.\; x = b]$'' \\
 \textbf{nitpick} \\[2\smallskipamount]
@@ -1691,7 +1722,7 @@
 of elements stored in the tree:
 
 \prew
-\textbf{theorem}~\textit{dataset\_skew\_split}:\\
+\textbf{theorem}~\textit{dataset\_skew\_split\/}:\\
 ``$\textit{dataset}~(\textit{skew}~t) = \textit{dataset}~t$'' \\
 ``$\textit{dataset}~(\textit{split}~t) = \textit{dataset}~t$'' \\
 \textbf{nitpick} \\[2\smallskipamount]
@@ -1702,7 +1733,7 @@
 should not alter the tree:
 
 \prew
-\textbf{theorem}~\textit{wf\_skew\_split}:\\
+\textbf{theorem}~\textit{wf\_skew\_split\/}:\\
 ``$\textit{wf}~t\,\Longrightarrow\, \textit{skew}~t = t$'' \\
 ``$\textit{wf}~t\,\Longrightarrow\, \textit{split}~t = t$'' \\
 \textbf{nitpick} \\[2\smallskipamount]
@@ -1723,7 +1754,7 @@
 \textit{split}. Let's see if this causes any problems:
 
 \prew
-\textbf{theorem}~\textit{wf\_insort}:\kern.4em ``$\textit{wf}~t\,\Longrightarrow\, \textit{wf}~(\textit{insort}~t~x)$'' \\
+\textbf{theorem}~\textit{wf\_insort\/}:\kern.4em ``$\textit{wf}~t\,\Longrightarrow\, \textit{wf}~(\textit{insort}~t~x)$'' \\
 \textbf{nitpick} \\[2\smallskipamount]
 \slshape Nitpick found a counterexample for \textit{card} $'a$ = 4: \\[2\smallskipamount]
 \hbox{}\qquad Free variables: \nopagebreak \\
@@ -1738,7 +1769,7 @@
 $\textit{insort}~t~x$ using the \textit{eval} option. This gives
 
 \prew
-\textbf{theorem} \textit{wf\_insort\_nat}:\kern.4em ``$\textit{wf}~t\,\Longrightarrow\, \textit{wf}~(\textit{insort}~t~(x\Colon\textit{nat}))$'' \\
+\textbf{theorem} \textit{wf\_insort\_nat\/}:\kern.4em ``$\textit{wf}~t\,\Longrightarrow\, \textit{wf}~(\textit{insort}~t~(x\Colon\textit{nat}))$'' \\
 \textbf{nitpick} [\textit{eval} = ``$\textit{insort}~t~x$''] \\[2\smallskipamount]
 \slshape Nitpick found a counterexample: \\[2\smallskipamount]
 \hbox{}\qquad Free variables: \nopagebreak \\
@@ -1755,7 +1786,7 @@
 Reintroducing the code seems to solve the problem:
 
 \prew
-\textbf{theorem}~\textit{wf\_insort}:\kern.4em ``$\textit{wf}~t\,\Longrightarrow\, \textit{wf}~(\textit{insort}~t~x)$'' \\
+\textbf{theorem}~\textit{wf\_insort\/}:\kern.4em ``$\textit{wf}~t\,\Longrightarrow\, \textit{wf}~(\textit{insort}~t~x)$'' \\
 \textbf{nitpick} \\[2\smallskipamount]
 {\slshape Nitpick ran out of time after checking 7 of 8 scopes.}
 \postw
@@ -1764,7 +1795,7 @@
 obvious way:
 
 \prew
-\textbf{theorem} \textit{dataset\_insort}:\kern.4em
+\textbf{theorem} \textit{dataset\_insort\/}:\kern.4em
 ``$\textit{dataset}~(\textit{insort}~t~x) = \{x\} \cup \textit{dataset}~t$'' \\
 \textbf{nitpick} \\[2\smallskipamount]
 {\slshape Nitpick ran out of time after checking 6 of 8 scopes.}
@@ -1825,19 +1856,19 @@
 
 \begin{enum}
 \item[$\bullet$] \qtybf{string}: A string.
-\item[$\bullet$] \qtybf{bool}: \textit{true} or \textit{false}.
-\item[$\bullet$] \qtybf{bool\_or\_smart}: \textit{true}, \textit{false}, or \textit{smart}.
-\item[$\bullet$] \qtybf{int}: An integer. Negative integers are prefixed with a hyphen.
-\item[$\bullet$] \qtybf{int\_or\_smart}: An integer or \textit{smart}.
+\item[$\bullet$] \qtybf{bool\/}: \textit{true} or \textit{false}.
+\item[$\bullet$] \qtybf{bool\_or\_smart\/}: \textit{true}, \textit{false}, or \textit{smart}.
+\item[$\bullet$] \qtybf{int\/}: An integer. Negative integers are prefixed with a hyphen.
+\item[$\bullet$] \qtybf{int\_or\_smart\/}: An integer or \textit{smart}.
 \item[$\bullet$] \qtybf{int\_range}: An integer (e.g., 3) or a range
 of nonnegative integers (e.g., $1$--$4$). The range symbol `--' can be entered as \texttt{-} (hyphen) or \texttt{\char`\\\char`\<midarrow\char`\>}.
 
 \item[$\bullet$] \qtybf{int\_seq}: A comma-separated sequence of ranges of integers (e.g.,~1{,}3{,}\allowbreak6--8).
 \item[$\bullet$] \qtybf{time}: An integer followed by $\textit{min}$ (minutes), $s$ (seconds), or \textit{ms}
 (milliseconds), or the keyword \textit{none} ($\infty$ years).
-\item[$\bullet$] \qtybf{const}: The name of a HOL constant.
+\item[$\bullet$] \qtybf{const\/}: The name of a HOL constant.
 \item[$\bullet$] \qtybf{term}: A HOL term (e.g., ``$f~x$'').
-\item[$\bullet$] \qtybf{term\_list}: A space-separated list of HOL terms (e.g.,
+\item[$\bullet$] \qtybf{term\_list\/}: A space-separated list of HOL terms (e.g.,
 ``$f~x$''~``$g~y$'').
 \item[$\bullet$] \qtybf{type}: A HOL type.
 \end{enum}
@@ -2190,7 +2221,6 @@
 the \textit{format}~\qty{term} option described above.
 \end{enum}
 
-%% MARK: Authentication
 \subsection{Authentication}
 \label{authentication}
 
@@ -2564,14 +2594,14 @@
 definition as follows:
 
 \prew
-\textbf{lemma} $\mathit{odd\_def}'$ [\textit{nitpick\_def}]: ``$\textit{odd}~n \,\equiv\, n~\textrm{mod}~2 = 1$''
+\textbf{lemma} $\mathit{odd\_def}'$ [\textit{nitpick\_def}]:\kern.4em ``$\textit{odd}~n \,\equiv\, n~\textrm{mod}~2 = 1$''
 \postw
 
 Nitpick then expands all occurrences of $\mathit{odd}~n$ to $n~\textrm{mod}~2
 = 1$. Alternatively, we can specify an equational specification of the constant:
 
 \prew
-\textbf{lemma} $\mathit{odd\_simp}'$ [\textit{nitpick\_simp}]: ``$\textit{odd}~n = (n~\textrm{mod}~2 = 1)$''
+\textbf{lemma} $\mathit{odd\_simp}'$ [\textit{nitpick\_simp}]:\kern.4em ``$\textit{odd}~n = (n~\textrm{mod}~2 = 1)$''
 \postw
 
 Such tweaks should be done with great care, because Nitpick will assume that the
--- a/doc-src/more_antiquote.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/doc-src/more_antiquote.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -91,9 +91,9 @@
     val (_, eqngr) = Code_Preproc.obtain thy [const] [];
     fun holize thm = @{thm meta_eq_to_obj_eq} OF [thm];
     val thms = Code_Preproc.cert eqngr const
-      |> Code.equations_thms_cert thy
+      |> Code.equations_of_cert thy
       |> snd
-      |> map_filter (fn (_, (thm, proper)) => if proper then SOME thm else NONE)
+      |> map_filter (fn (_, (some_thm, proper)) => if proper then some_thm else NONE)
       |> map (holize o no_vars ctxt o AxClass.overload thy);
   in ThyOutput.output (ThyOutput.maybe_pretty_source (pretty_thm ctxt) src thms) end;
 
--- a/etc/isar-keywords-ZF.el	Mon Feb 22 11:17:41 2010 -0800
+++ b/etc/isar-keywords-ZF.el	Mon Feb 22 11:19:15 2010 -0800
@@ -44,6 +44,7 @@
     "classes"
     "classrel"
     "codatatype"
+    "code_abstype"
     "code_datatype"
     "code_library"
     "code_module"
@@ -413,7 +414,8 @@
   '("inductive_cases"))
 
 (defconst isar-keywords-theory-goal
-  '("corollary"
+  '("code_abstype"
+    "corollary"
     "instance"
     "interpretation"
     "lemma"
--- a/etc/isar-keywords.el	Mon Feb 22 11:17:41 2010 -0800
+++ b/etc/isar-keywords.el	Mon Feb 22 11:19:15 2010 -0800
@@ -188,6 +188,9 @@
     "print_locales"
     "print_methods"
     "print_orders"
+    "print_quotconsts"
+    "print_quotients"
+    "print_quotmaps"
     "print_rules"
     "print_simpset"
     "print_statement"
@@ -203,6 +206,8 @@
     "quickcheck"
     "quickcheck_params"
     "quit"
+    "quotient_definition"
+    "quotient_type"
     "realizability"
     "realizers"
     "recdef"
@@ -395,6 +400,9 @@
     "print_locales"
     "print_methods"
     "print_orders"
+    "print_quotconsts"
+    "print_quotients"
+    "print_quotmaps"
     "print_rules"
     "print_simpset"
     "print_statement"
@@ -512,6 +520,7 @@
     "print_ast_translation"
     "print_translation"
     "quickcheck_params"
+    "quotient_definition"
     "realizability"
     "realizers"
     "recdef"
@@ -550,6 +559,7 @@
     "nominal_inductive2"
     "nominal_primrec"
     "pcpodef"
+    "quotient_type"
     "recdef_tc"
     "rep_datatype"
     "specification"
--- a/src/Cube/Cube.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Cube/Cube.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -18,43 +18,43 @@
   context' typing'
 
 consts
-  Abs           :: "[term, term => term] => term"
-  Prod          :: "[term, term => term] => term"
-  Trueprop      :: "[context, typing] => prop"
-  MT_context    :: "context"
-  Context       :: "[typing, context] => context"
-  star          :: "term"                               ("*")
-  box           :: "term"                               ("[]")
-  app           :: "[term, term] => term"               (infixl "^" 20)
-  Has_type      :: "[term, term] => typing"
+  Abs :: "[term, term => term] => term"
+  Prod :: "[term, term => term] => term"
+  Trueprop :: "[context, typing] => prop"
+  MT_context :: "context"
+  Context :: "[typing, context] => context"
+  star :: "term"    ("*")
+  box :: "term"    ("[]")
+  app :: "[term, term] => term"    (infixl "^" 20)
+  Has_type :: "[term, term] => typing"
 
 syntax
-  Trueprop      :: "[context', typing'] => prop"        ("(_/ |- _)")
-  Trueprop1     :: "typing' => prop"                    ("(_)")
-  ""            :: "id => context'"                     ("_")
-  ""            :: "var => context'"                    ("_")
-  MT_context    :: "context'"                           ("")
-  Context       :: "[typing', context'] => context'"    ("_ _")
-  Has_type      :: "[term, term] => typing'"            ("(_:/ _)" [0, 0] 5)
-  Lam           :: "[idt, term, term] => term"          ("(3Lam _:_./ _)" [0, 0, 0] 10)
-  Pi            :: "[idt, term, term] => term"          ("(3Pi _:_./ _)" [0, 0] 10)
-  arrow         :: "[term, term] => term"               (infixr "->" 10)
+  "\<^const>Cube.Trueprop" :: "[context', typing'] => prop"    ("(_/ |- _)")
+  "_Trueprop1" :: "typing' => prop"    ("(_)")
+  "" :: "id => context'"    ("_")
+  "" :: "var => context'"    ("_")
+  "\<^const>Cube.MT_context" :: "context'"    ("")
+  "\<^const>Cube.Context" :: "[typing', context'] => context'"    ("_ _")
+  "\<^const>Cube.Has_type" :: "[term, term] => typing'"    ("(_:/ _)" [0, 0] 5)
+  "_Lam" :: "[idt, term, term] => term"    ("(3Lam _:_./ _)" [0, 0, 0] 10)
+  "_Pi" :: "[idt, term, term] => term"    ("(3Pi _:_./ _)" [0, 0] 10)
+  "_arrow" :: "[term, term] => term"    (infixr "->" 10)
 
 translations
   ("prop") "x:X" == ("prop") "|- x:X"
-  "Lam x:A. B"   == "CONST Abs(A, %x. B)"
-  "Pi x:A. B"    => "CONST Prod(A, %x. B)"
-  "A -> B"       => "CONST Prod(A, %_. B)"
+  "Lam x:A. B" == "CONST Abs(A, %x. B)"
+  "Pi x:A. B" => "CONST Prod(A, %x. B)"
+  "A -> B" => "CONST Prod(A, %_. B)"
 
 syntax (xsymbols)
-  Trueprop      :: "[context', typing'] => prop"        ("(_/ \<turnstile> _)")
-  box           :: "term"                               ("\<box>")
-  Lam           :: "[idt, term, term] => term"          ("(3\<Lambda> _:_./ _)" [0, 0, 0] 10)
-  Pi            :: "[idt, term, term] => term"          ("(3\<Pi> _:_./ _)" [0, 0] 10)
-  arrow         :: "[term, term] => term"               (infixr "\<rightarrow>" 10)
+  "\<^const>Cube.Trueprop" :: "[context', typing'] => prop"    ("(_/ \<turnstile> _)")
+  "\<^const>Cube.box" :: "term"    ("\<box>")
+  "_Lam" :: "[idt, term, term] => term"    ("(3\<Lambda> _:_./ _)" [0, 0, 0] 10)
+  "_Pi" :: "[idt, term, term] => term"    ("(3\<Pi> _:_./ _)" [0, 0] 10)
+  "_arrow" :: "[term, term] => term"    (infixr "\<rightarrow>" 10)
 
 print_translation {*
-  [(@{const_syntax Prod}, dependent_tr' (@{syntax_const Pi}, @{syntax_const arrow}))]
+  [(@{const_syntax Prod}, dependent_tr' (@{syntax_const "_Pi"}, @{syntax_const "_arrow"}))]
 *}
 
 axioms
--- a/src/HOL/Algebra/Divisibility.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Algebra/Divisibility.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -2250,7 +2250,7 @@
   assumes ab: "a divides b"
     and afs: "wfactors G as a" and bfs: "wfactors G bs b"
     and carr: "a \<in> carrier G"  "b \<in> carrier G"  "set as \<subseteq> carrier G"  "set bs \<subseteq> carrier G"
-  shows "fmset G as \<le># fmset G bs"
+  shows "fmset G as \<le> fmset G bs"
 using ab
 proof (elim dividesE)
   fix c
@@ -2270,7 +2270,7 @@
 qed
 
 lemma (in comm_monoid_cancel) fmsubset_divides:
-  assumes msubset: "fmset G as \<le># fmset G bs"
+  assumes msubset: "fmset G as \<le> fmset G bs"
     and afs: "wfactors G as a" and bfs: "wfactors G bs b"
     and acarr: "a \<in> carrier G" and bcarr: "b \<in> carrier G"
     and ascarr: "set as \<subseteq> carrier G" and bscarr: "set bs \<subseteq> carrier G"
@@ -2323,7 +2323,7 @@
   assumes "wfactors G as a" and "wfactors G bs b"
     and "a \<in> carrier G" and "b \<in> carrier G" 
     and "set as \<subseteq> carrier G" and "set bs \<subseteq> carrier G"
-  shows "a divides b = (fmset G as \<le># fmset G bs)"
+  shows "a divides b = (fmset G as \<le> fmset G bs)"
 using assms
 by (blast intro: divides_fmsubset fmsubset_divides)
 
@@ -2331,7 +2331,7 @@
 text {* Proper factors on multisets *}
 
 lemma (in factorial_monoid) fmset_properfactor:
-  assumes asubb: "fmset G as \<le># fmset G bs"
+  assumes asubb: "fmset G as \<le> fmset G bs"
     and anb: "fmset G as \<noteq> fmset G bs"
     and "wfactors G as a" and "wfactors G bs b"
     and "a \<in> carrier G" and "b \<in> carrier G"
@@ -2341,10 +2341,10 @@
 apply (rule fmsubset_divides[of as bs], fact+)
 proof
   assume "b divides a"
-  hence "fmset G bs \<le># fmset G as"
+  hence "fmset G bs \<le> fmset G as"
       by (rule divides_fmsubset) fact+
   with asubb
-      have "fmset G as = fmset G bs" by (simp add: mset_le_antisym)
+      have "fmset G as = fmset G bs" by (rule order_antisym)
   with anb
       show "False" ..
 qed
@@ -2354,7 +2354,7 @@
     and "wfactors G as a" and "wfactors G bs b"
     and "a \<in> carrier G" and "b \<in> carrier G"
     and "set as \<subseteq> carrier G" and "set bs \<subseteq> carrier G"
-  shows "fmset G as \<le># fmset G bs \<and> fmset G as \<noteq> fmset G bs"
+  shows "fmset G as \<le> fmset G bs \<and> fmset G as \<noteq> fmset G bs"
 using pf
 apply (elim properfactorE)
 apply rule
@@ -2738,12 +2738,12 @@
   have "c gcdof a b"
   proof (simp add: isgcd_def, safe)
     from csmset
-        have "fmset G cs \<le># fmset G as"
+        have "fmset G cs \<le> fmset G as"
         by (simp add: multiset_inter_def mset_le_def)
     thus "c divides a" by (rule fmsubset_divides) fact+
   next
     from csmset
-        have "fmset G cs \<le># fmset G bs"
+        have "fmset G cs \<le> fmset G bs"
         by (simp add: multiset_inter_def mset_le_def, force)
     thus "c divides b" by (rule fmsubset_divides) fact+
   next
@@ -2756,13 +2756,13 @@
         by auto
 
     assume "y divides a"
-    hence ya: "fmset G ys \<le># fmset G as" by (rule divides_fmsubset) fact+
+    hence ya: "fmset G ys \<le> fmset G as" by (rule divides_fmsubset) fact+
 
     assume "y divides b"
-    hence yb: "fmset G ys \<le># fmset G bs" by (rule divides_fmsubset) fact+
+    hence yb: "fmset G ys \<le> fmset G bs" by (rule divides_fmsubset) fact+
 
     from ya yb csmset
-    have "fmset G ys \<le># fmset G cs" by (simp add: mset_le_def multiset_inter_count)
+    have "fmset G ys \<le> fmset G cs" by (simp add: mset_le_def multiset_inter_count)
     thus "y divides c" by (rule fmsubset_divides) fact+
   qed
 
@@ -2837,10 +2837,10 @@
 
   have "c lcmof a b"
   proof (simp add: islcm_def, safe)
-    from csmset have "fmset G as \<le># fmset G cs" by (simp add: mset_le_def, force)
+    from csmset have "fmset G as \<le> fmset G cs" by (simp add: mset_le_def, force)
     thus "a divides c" by (rule fmsubset_divides) fact+
   next
-    from csmset have "fmset G bs \<le># fmset G cs" by (simp add: mset_le_def)
+    from csmset have "fmset G bs \<le> fmset G cs" by (simp add: mset_le_def)
     thus "b divides c" by (rule fmsubset_divides) fact+
   next
     fix y
@@ -2852,13 +2852,13 @@
         by auto
 
     assume "a divides y"
-    hence ya: "fmset G as \<le># fmset G ys" by (rule divides_fmsubset) fact+
+    hence ya: "fmset G as \<le> fmset G ys" by (rule divides_fmsubset) fact+
 
     assume "b divides y"
-    hence yb: "fmset G bs \<le># fmset G ys" by (rule divides_fmsubset) fact+
+    hence yb: "fmset G bs \<le> fmset G ys" by (rule divides_fmsubset) fact+
 
     from ya yb csmset
-    have "fmset G cs \<le># fmset G ys"
+    have "fmset G cs \<le> fmset G ys"
       apply (simp add: mset_le_def, clarify)
       apply (case_tac "count (fmset G as) a < count (fmset G bs) a")
        apply simp
--- a/src/HOL/Algebra/abstract/Ring2.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Algebra/abstract/Ring2.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -205,8 +205,8 @@
 ML {*
 fun ring_ord (Const (a, _)) =
     find_index (fn a' => a = a')
-      [@{const_name Algebras.zero}, @{const_name Algebras.plus}, @{const_name Algebras.uminus},
-        @{const_name Algebras.minus}, @{const_name Algebras.one}, @{const_name Algebras.times}]
+      [@{const_name Groups.zero}, @{const_name Groups.plus}, @{const_name Groups.uminus},
+        @{const_name Groups.minus}, @{const_name Groups.one}, @{const_name Groups.times}]
   | ring_ord _ = ~1;
 
 fun termless_ring (a, b) = (TermOrd.term_lpo ring_ord (a, b) = LESS);
--- a/src/HOL/Algebras.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Algebras.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -8,8 +8,6 @@
 imports HOL
 begin
 
-subsection {* Generic algebraic structures *}
-
 text {*
   These locales provide basic structures for interpretation into
   bigger structures;  extensions require careful thinking, otherwise
@@ -54,58 +52,4 @@
 
 end
 
-
-subsection {* Generic syntactic operations *}
-
-class zero = 
-  fixes zero :: 'a  ("0")
-
-class one =
-  fixes one  :: 'a  ("1")
-
-hide (open) const zero one
-
-syntax
-  "_index1"  :: index    ("\<^sub>1")
-translations
-  (index) "\<^sub>1" => (index) "\<^bsub>\<struct>\<^esub>"
-
-lemma Let_0 [simp]: "Let 0 f = f 0"
-  unfolding Let_def ..
-
-lemma Let_1 [simp]: "Let 1 f = f 1"
-  unfolding Let_def ..
-
-setup {*
-  Reorient_Proc.add
-    (fn Const(@{const_name Algebras.zero}, _) => true
-      | Const(@{const_name Algebras.one}, _) => true
-      | _ => false)
-*}
-
-simproc_setup reorient_zero ("0 = x") = Reorient_Proc.proc
-simproc_setup reorient_one ("1 = x") = Reorient_Proc.proc
-
-typed_print_translation {*
-let
-  fun tr' c = (c, fn show_sorts => fn T => fn ts =>
-    if (not o null) ts orelse T = dummyT
-      orelse not (! show_types) andalso can Term.dest_Type T
-    then raise Match
-    else Syntax.const Syntax.constrainC $ Syntax.const c $ Syntax.term_of_typ show_sorts T);
-in map tr' [@{const_syntax Algebras.one}, @{const_syntax Algebras.zero}] end;
-*} -- {* show types that are presumably too general *}
-
-class plus =
-  fixes plus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "+" 65)
-
-class minus =
-  fixes minus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "-" 65)
-
-class uminus =
-  fixes uminus :: "'a \<Rightarrow> 'a"  ("- _" [81] 80)
-
-class times =
-  fixes times :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "*" 70)
-
 end
\ No newline at end of file
--- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -670,13 +670,13 @@
  end
 
 fun whatis x ct = case term_of ct of
-  Const(@{const_name Algebras.plus}, _)$(Const(@{const_name Algebras.times},_)$_$y)$_ =>
+  Const(@{const_name Groups.plus}, _)$(Const(@{const_name Groups.times},_)$_$y)$_ =>
      if y aconv term_of x then ("c*x+t",[(funpow 2 Thm.dest_arg1) ct, Thm.dest_arg ct])
      else ("Nox",[])
-| Const(@{const_name Algebras.plus}, _)$y$_ =>
+| Const(@{const_name Groups.plus}, _)$y$_ =>
      if y aconv term_of x then ("x+t",[Thm.dest_arg ct])
      else ("Nox",[])
-| Const(@{const_name Algebras.times}, _)$_$y =>
+| Const(@{const_name Groups.times}, _)$_$y =>
      if y aconv term_of x then ("c*x",[Thm.dest_arg1 ct])
      else ("Nox",[])
 | t => if t aconv term_of x then ("x",[]) else ("Nox",[]);
@@ -684,7 +684,7 @@
 fun xnormalize_conv ctxt [] ct = reflexive ct
 | xnormalize_conv ctxt (vs as (x::_)) ct =
    case term_of ct of
-   Const(@{const_name Orderings.less},_)$_$Const(@{const_name Algebras.zero},_) =>
+   Const(@{const_name Orderings.less},_)$_$Const(@{const_name Groups.zero},_) =>
     (case whatis x (Thm.dest_arg1 ct) of
     ("c*x+t",[c,t]) =>
        let
@@ -727,7 +727,7 @@
     | _ => reflexive ct)
 
 
-|  Const(@{const_name Orderings.less_eq},_)$_$Const(@{const_name Algebras.zero},_) =>
+|  Const(@{const_name Orderings.less_eq},_)$_$Const(@{const_name Groups.zero},_) =>
    (case whatis x (Thm.dest_arg1 ct) of
     ("c*x+t",[c,t]) =>
        let
@@ -771,7 +771,7 @@
       in rth end
     | _ => reflexive ct)
 
-|  Const("op =",_)$_$Const(@{const_name Algebras.zero},_) =>
+|  Const("op =",_)$_$Const(@{const_name Groups.zero},_) =>
    (case whatis x (Thm.dest_arg1 ct) of
     ("c*x+t",[c,t]) =>
        let
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -2947,10 +2947,10 @@
 fun rrelT rT = [rT,rT] ---> rT;
 fun rrT rT = [rT, rT] ---> bT;
 fun divt rT = Const(@{const_name Rings.divide},rrelT rT);
-fun timest rT = Const(@{const_name Algebras.times},rrelT rT);
-fun plust rT = Const(@{const_name Algebras.plus},rrelT rT);
-fun minust rT = Const(@{const_name Algebras.minus},rrelT rT);
-fun uminust rT = Const(@{const_name Algebras.uminus}, rT --> rT);
+fun timest rT = Const(@{const_name Groups.times},rrelT rT);
+fun plust rT = Const(@{const_name Groups.plus},rrelT rT);
+fun minust rT = Const(@{const_name Groups.minus},rrelT rT);
+fun uminust rT = Const(@{const_name Groups.uminus}, rT --> rT);
 fun powt rT = Const(@{const_name "power"}, [rT,@{typ "nat"}] ---> rT);
 val brT = [bT, bT] ---> bT;
 val nott = @{term "Not"};
@@ -2961,7 +2961,7 @@
 fun llt rT = Const(@{const_name Orderings.less},rrT rT);
 fun lle rT = Const(@{const_name Orderings.less},rrT rT);
 fun eqt rT = Const("op =",rrT rT);
-fun rz rT = Const(@{const_name Algebras.zero},rT);
+fun rz rT = Const(@{const_name Groups.zero},rT);
 
 fun dest_nat t = case t of
   Const ("Suc",_)$t' => 1 + dest_nat t'
@@ -2969,10 +2969,10 @@
 
 fun num_of_term m t = 
  case t of
-   Const(@{const_name Algebras.uminus},_)$t => @{code poly.Neg} (num_of_term m t)
- | Const(@{const_name Algebras.plus},_)$a$b => @{code poly.Add} (num_of_term m a, num_of_term m b)
- | Const(@{const_name Algebras.minus},_)$a$b => @{code poly.Sub} (num_of_term m a, num_of_term m b)
- | Const(@{const_name Algebras.times},_)$a$b => @{code poly.Mul} (num_of_term m a, num_of_term m b)
+   Const(@{const_name Groups.uminus},_)$t => @{code poly.Neg} (num_of_term m t)
+ | Const(@{const_name Groups.plus},_)$a$b => @{code poly.Add} (num_of_term m a, num_of_term m b)
+ | Const(@{const_name Groups.minus},_)$a$b => @{code poly.Sub} (num_of_term m a, num_of_term m b)
+ | Const(@{const_name Groups.times},_)$a$b => @{code poly.Mul} (num_of_term m a, num_of_term m b)
  | Const(@{const_name Power.power},_)$a$n => @{code poly.Pw} (num_of_term m a, dest_nat n)
  | Const(@{const_name Rings.divide},_)$a$b => @{code poly.C} (HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
  | _ => (@{code poly.C} (HOLogic.dest_number t |> snd,1) 
@@ -2980,10 +2980,10 @@
 
 fun tm_of_term m m' t = 
  case t of
-   Const(@{const_name Algebras.uminus},_)$t => @{code Neg} (tm_of_term m m' t)
- | Const(@{const_name Algebras.plus},_)$a$b => @{code Add} (tm_of_term m m' a, tm_of_term m m' b)
- | Const(@{const_name Algebras.minus},_)$a$b => @{code Sub} (tm_of_term m m' a, tm_of_term m m' b)
- | Const(@{const_name Algebras.times},_)$a$b => @{code Mul} (num_of_term m' a, tm_of_term m m' b)
+   Const(@{const_name Groups.uminus},_)$t => @{code Neg} (tm_of_term m m' t)
+ | Const(@{const_name Groups.plus},_)$a$b => @{code Add} (tm_of_term m m' a, tm_of_term m m' b)
+ | Const(@{const_name Groups.minus},_)$a$b => @{code Sub} (tm_of_term m m' a, tm_of_term m m' b)
+ | Const(@{const_name Groups.times},_)$a$b => @{code Mul} (num_of_term m' a, tm_of_term m m' b)
  | _ => (@{code CP} (num_of_term m' t) 
          handle TERM _ => @{code Bound} (AList.lookup (op aconv) m t |> the)
               | Option => @{code Bound} (AList.lookup (op aconv) m t |> the));
--- a/src/HOL/Decision_Procs/commutative_ring_tac.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Decision_Procs/commutative_ring_tac.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -13,8 +13,8 @@
 struct
 
 (* Zero and One of the commutative ring *)
-fun cring_zero T = Const (@{const_name Algebras.zero}, T);
-fun cring_one T = Const (@{const_name Algebras.one}, T);
+fun cring_zero T = Const (@{const_name Groups.zero}, T);
+fun cring_one T = Const (@{const_name Groups.one}, T);
 
 (* reification functions *)
 (* add two polynom expressions *)
@@ -49,13 +49,13 @@
   | reif_pol T vs t = pol_Pc T $ t;
 
 (* reification of polynom expressions *)
-fun reif_polex T vs (Const (@{const_name Algebras.plus}, _) $ a $ b) =
+fun reif_polex T vs (Const (@{const_name Groups.plus}, _) $ a $ b) =
       polex_add T $ reif_polex T vs a $ reif_polex T vs b
-  | reif_polex T vs (Const (@{const_name Algebras.minus}, _) $ a $ b) =
+  | reif_polex T vs (Const (@{const_name Groups.minus}, _) $ a $ b) =
       polex_sub T $ reif_polex T vs a $ reif_polex T vs b
-  | reif_polex T vs (Const (@{const_name Algebras.times}, _) $ a $ b) =
+  | reif_polex T vs (Const (@{const_name Groups.times}, _) $ a $ b) =
       polex_mul T $ reif_polex T vs a $ reif_polex T vs b
-  | reif_polex T vs (Const (@{const_name Algebras.uminus}, _) $ a) =
+  | reif_polex T vs (Const (@{const_name Groups.uminus}, _) $ a) =
       polex_neg T $ reif_polex T vs a
   | reif_polex T vs (Const (@{const_name Power.power}, _) $ a $ n) =
       polex_pow T $ reif_polex T vs a $ n
--- a/src/HOL/Finite_Set.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Finite_Set.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -1055,7 +1055,7 @@
 
 subsection {* Generalized summation over a set *}
 
-interpretation comm_monoid_add: comm_monoid_mult "0::'a::comm_monoid_add" "op +"
+interpretation comm_monoid_add: comm_monoid_mult "op +" "0::'a::comm_monoid_add"
   proof qed (auto intro: add_assoc add_commute)
 
 definition setsum :: "('a => 'b) => 'a set => 'b::comm_monoid_add"
--- a/src/HOL/Groups.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Groups.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -6,9 +6,64 @@
 
 theory Groups
 imports Orderings
-uses "~~/src/Provers/Arith/abel_cancel.ML"
+uses ("~~/src/Provers/Arith/abel_cancel.ML")
 begin
 
+subsection {* Generic operations *}
+
+class zero = 
+  fixes zero :: 'a  ("0")
+
+class one =
+  fixes one  :: 'a  ("1")
+
+hide (open) const zero one
+
+syntax
+  "_index1"  :: index    ("\<^sub>1")
+translations
+  (index) "\<^sub>1" => (index) "\<^bsub>\<struct>\<^esub>"
+
+lemma Let_0 [simp]: "Let 0 f = f 0"
+  unfolding Let_def ..
+
+lemma Let_1 [simp]: "Let 1 f = f 1"
+  unfolding Let_def ..
+
+setup {*
+  Reorient_Proc.add
+    (fn Const(@{const_name Groups.zero}, _) => true
+      | Const(@{const_name Groups.one}, _) => true
+      | _ => false)
+*}
+
+simproc_setup reorient_zero ("0 = x") = Reorient_Proc.proc
+simproc_setup reorient_one ("1 = x") = Reorient_Proc.proc
+
+typed_print_translation {*
+let
+  fun tr' c = (c, fn show_sorts => fn T => fn ts =>
+    if (not o null) ts orelse T = dummyT
+      orelse not (! show_types) andalso can Term.dest_Type T
+    then raise Match
+    else Syntax.const Syntax.constrainC $ Syntax.const c $ Syntax.term_of_typ show_sorts T);
+in map tr' [@{const_syntax Groups.one}, @{const_syntax Groups.zero}] end;
+*} -- {* show types that are presumably too general *}
+
+class plus =
+  fixes plus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "+" 65)
+
+class minus =
+  fixes minus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "-" 65)
+
+class uminus =
+  fixes uminus :: "'a \<Rightarrow> 'a"  ("- _" [81] 80)
+
+class times =
+  fixes times :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "*" 70)
+
+use "~~/src/Provers/Arith/abel_cancel.ML"
+
 text {*
   The theory of partially ordered groups is taken from the books:
   \begin{itemize}
@@ -1129,8 +1184,8 @@
 (* term order for abelian groups *)
 
 fun agrp_ord (Const (a, _)) = find_index (fn a' => a = a')
-      [@{const_name Algebras.zero}, @{const_name Algebras.plus},
-        @{const_name Algebras.uminus}, @{const_name Algebras.minus}]
+      [@{const_name Groups.zero}, @{const_name Groups.plus},
+        @{const_name Groups.uminus}, @{const_name Groups.minus}]
   | agrp_ord _ = ~1;
 
 fun termless_agrp (a, b) = (TermOrd.term_lpo agrp_ord (a, b) = LESS);
@@ -1139,9 +1194,9 @@
   val ac1 = mk_meta_eq @{thm add_assoc};
   val ac2 = mk_meta_eq @{thm add_commute};
   val ac3 = mk_meta_eq @{thm add_left_commute};
-  fun solve_add_ac thy _ (_ $ (Const (@{const_name Algebras.plus},_) $ _ $ _) $ _) =
+  fun solve_add_ac thy _ (_ $ (Const (@{const_name Groups.plus},_) $ _ $ _) $ _) =
         SOME ac1
-    | solve_add_ac thy _ (_ $ x $ (Const (@{const_name Algebras.plus},_) $ y $ z)) =
+    | solve_add_ac thy _ (_ $ x $ (Const (@{const_name Groups.plus},_) $ y $ z)) =
         if termless_agrp (y, x) then SOME ac3 else NONE
     | solve_add_ac thy _ (_ $ x $ y) =
         if termless_agrp (y, x) then SOME ac2 else NONE
--- a/src/HOL/Import/Generate-HOL/GenHOL4Base.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Import/Generate-HOL/GenHOL4Base.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -182,9 +182,9 @@
   ">="         > HOL4Compat.nat_ge
   FUNPOW       > HOL4Compat.FUNPOW
   "<="         > Orderings.less_eq :: "[nat,nat]=>bool"
-  "+"          > Algebras.plus :: "[nat,nat]=>nat"
-  "*"          > Algebras.times :: "[nat,nat]=>nat"
-  "-"          > Algebras.minus :: "[nat,nat]=>nat"
+  "+"          > Groups.plus :: "[nat,nat]=>nat"
+  "*"          > Groups.times :: "[nat,nat]=>nat"
+  "-"          > Groups.minus :: "[nat,nat]=>nat"
   MIN          > Orderings.min :: "[nat,nat]=>nat"
   MAX          > Orderings.max :: "[nat,nat]=>nat"
   DIV          > Divides.div :: "[nat,nat]=>nat"
--- a/src/HOL/Import/Generate-HOL/GenHOL4Real.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Import/Generate-HOL/GenHOL4Real.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -16,12 +16,12 @@
   real > RealDef.real;
 
 const_maps
-  real_0   > Algebras.zero      :: real
-  real_1   > Algebras.one       :: real
-  real_neg > Algebras.uminus    :: "real => real"
-  inv      > Algebras.inverse   :: "real => real"
-  real_add > Algebras.plus      :: "[real,real] => real"
-  real_mul > Algebras.times     :: "[real,real] => real"
+  real_0   > Groups.zero      :: real
+  real_1   > Groups.one       :: real
+  real_neg > Groups.uminus    :: "real => real"
+  inv      > Groups.inverse   :: "real => real"
+  real_add > Groups.plus      :: "[real,real] => real"
+  real_mul > Groups.times     :: "[real,real] => real"
   real_lt  > Orderings.less      :: "[real,real] => bool";
 
 ignore_thms
@@ -51,8 +51,8 @@
   real_gt     > HOL4Compat.real_gt
   real_ge     > HOL4Compat.real_ge
   real_lte    > Orderings.less_eq :: "[real,real] => bool"
-  real_sub    > Algebras.minus :: "[real,real] => real"
-  "/"         > Algebras.divide :: "[real,real] => real"
+  real_sub    > Groups.minus :: "[real,real] => real"
+  "/"         > Rings.divide :: "[real,real] => real"
   pow         > Power.power :: "[real,nat] => real"
   abs         > Groups.abs :: "real => real"
   real_of_num > RealDef.real :: "nat => real";
--- a/src/HOL/Import/Generate-HOLLight/GenHOLLight.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Import/Generate-HOLLight/GenHOLLight.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -76,9 +76,9 @@
   SUC > Suc
   PRE > HOLLightCompat.Pred
   NUMERAL > HOL4Compat.NUMERAL
-  "+" > Algebras.plus :: "nat \<Rightarrow> nat \<Rightarrow> nat" 
-  "*" > Algebras.times :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-  "-" > Algebras.minus :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+  "+" > Groups.plus :: "nat \<Rightarrow> nat \<Rightarrow> nat" 
+  "*" > Groups.times :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+  "-" > Groups.minus :: "nat \<Rightarrow> nat \<Rightarrow> nat"
   BIT0 > HOLLightCompat.NUMERAL_BIT0
   BIT1 > HOL4Compat.NUMERAL_BIT1
   INL > Sum_Type.Inl
--- a/src/HOL/Import/HOL/arithmetic.imp	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Import/HOL/arithmetic.imp	Mon Feb 22 11:19:15 2010 -0800
@@ -24,9 +24,9 @@
   ">=" > "HOL4Compat.nat_ge"
   ">" > "HOL4Compat.nat_gt"
   "<=" > "Orderings.less_eq" :: "nat => nat => bool"
-  "-" > "Algebras.minus_class.minus" :: "nat => nat => nat"
-  "+" > "Algebras.plus_class.plus" :: "nat => nat => nat"
-  "*" > "Algebras.times_class.times" :: "nat => nat => nat"
+  "-" > "Groups.minus" :: "nat => nat => nat"
+  "+" > "Groups.plus" :: "nat => nat => nat"
+  "*" > "Groups.times" :: "nat => nat => nat"
 
 thm_maps
   "num_case_def" > "HOL4Compat.num_case_def"
--- a/src/HOL/Import/HOL/real.imp	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Import/HOL/real.imp	Mon Feb 22 11:19:15 2010 -0800
@@ -10,7 +10,7 @@
 const_maps
   "sup" > "HOL4Real.real.sup"
   "sum" > "HOL4Real.real.sum"
-  "real_sub" > "Algebras.minus" :: "real => real => real"
+  "real_sub" > "Groups.minus" :: "real => real => real"
   "real_of_num" > "RealDef.real" :: "nat => real"
   "real_lte" > "Orderings.less_eq" :: "real => real => bool"
   "real_gt" > "HOL4Compat.real_gt"
--- a/src/HOL/Import/HOL/realax.imp	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Import/HOL/realax.imp	Mon Feb 22 11:19:15 2010 -0800
@@ -27,12 +27,12 @@
   "treal_add" > "HOL4Real.realax.treal_add"
   "treal_1" > "HOL4Real.realax.treal_1"
   "treal_0" > "HOL4Real.realax.treal_0"
-  "real_neg" > "Algebras.uminus_class.uminus" :: "real => real"
-  "real_mul" > "Algebras.times_class.times" :: "real => real => real"
+  "real_neg" > "Groups.uminus" :: "real => real"
+  "real_mul" > "Groups.times" :: "real => real => real"
   "real_lt" > "Orderings.less" :: "real => real => bool"
-  "real_add" > "Algebras.plus_class.plus" :: "real => real => real"
-  "real_1" > "Algebras.one_class.one" :: "real"
-  "real_0" > "Algebras.zero_class.zero" :: "real"
+  "real_add" > "Groups.plus" :: "real => real => real"
+  "real_1" > "Groups.one" :: "real"
+  "real_0" > "Groups.zero" :: "real"
   "inv" > "Ring.inverse" :: "real => real"
   "hreal_of_treal" > "HOL4Real.realax.hreal_of_treal"
 
--- a/src/HOL/Import/HOLLight/hollight.imp	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Import/HOLLight/hollight.imp	Mon Feb 22 11:19:15 2010 -0800
@@ -238,10 +238,10 @@
   "<=" > "HOLLight.hollight.<="
   "<" > "HOLLight.hollight.<"
   "/\\" > "op &"
-  "-" > "Algebras.minus_class.minus" :: "nat => nat => nat"
+  "-" > "Groups.minus" :: "nat => nat => nat"
   "," > "Pair"
-  "+" > "Algebras.plus_class.plus" :: "nat => nat => nat"
-  "*" > "Algebras.times_class.times" :: "nat => nat => nat"
+  "+" > "Groups.plus" :: "nat => nat => nat"
+  "*" > "Groups.times" :: "nat => nat => nat"
   "$" > "HOLLight.hollight.$"
   "!" > "All"
 
--- a/src/HOL/Import/hol4rews.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Import/hol4rews.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -613,13 +613,13 @@
     end
 
 local
-    fun handle_meta [x as Appl[Appl[Constant "_constrain", Constant "==", _],_,_]] = x
-      | handle_meta [x as Appl[Appl[Constant "_constrain", Constant "all", _],_]] = x
-      | handle_meta [x as Appl[Appl[Constant "_constrain", Constant "==>", _],_,_]] = x
+    fun handle_meta [x as Appl[Appl[Constant "_constrain", Constant @{const_syntax "=="}, _],_,_]] = x
+      | handle_meta [x as Appl[Appl[Constant "_constrain", Constant @{const_syntax all}, _],_]] = x
+      | handle_meta [x as Appl[Appl[Constant "_constrain", Constant @{const_syntax "==>"}, _],_,_]] = x
       | handle_meta [x] = Appl[Constant "Trueprop",x]
       | handle_meta _ = error "hol4rews error: Trueprop not applied to single argument"
 in
-val smarter_trueprop_parsing = [("Trueprop",handle_meta)]
+val smarter_trueprop_parsing = [(@{const_syntax Trueprop},handle_meta)]
 end
 
 local
--- a/src/HOL/Induct/ROOT.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Induct/ROOT.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -1,5 +1,5 @@
 setmp_noncritical quick_and_dirty true
   use_thys ["Common_Patterns"];
 
-use_thys ["QuoDataType", "QuoNestedDataType", "Term",
+use_thys ["QuoDataType", "QuoNestedDataType", "Term", "SList",
   "ABexp", "Tree", "Ordinals", "Sigma_Algebra", "Comb", "PropLog", "Com"];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Induct/SList.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -0,0 +1,434 @@
+(*  Title:      HOL/Induct/SList.thy
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+
+This theory is strictly of historical interest. It illustrates how
+recursive datatypes can be constructed in higher-order logic from
+first principles. Isabelle's datatype package automates a
+construction of this sort.
+
+Enriched theory of lists; mutual indirect recursive data-types.
+
+Definition of type 'a list (strict lists) by a least fixed point
+
+We use          list(A) == lfp(%Z. {NUMB(0)} <+> A <*> Z)
+and not         list    == lfp(%Z. {NUMB(0)} <+> range(Leaf) <*> Z)
+
+so that list can serve as a "functor" for defining other recursive types.
+
+This enables the conservative construction of mutual recursive data-types
+such as
+
+datatype 'a m = Node 'a * ('a m) list
+*)
+
+header {* Extended List Theory (old) *}
+
+theory SList
+imports Sexp
+begin
+
+(*Hilbert_Choice is needed for the function "inv"*)
+
+(* *********************************************************************** *)
+(*                                                                         *)
+(* Building up data type                                                   *)
+(*                                                                         *)
+(* *********************************************************************** *)
+
+
+(* Defining the Concrete Constructors *)
+definition
+  NIL  :: "'a item" where
+  "NIL = In0(Numb(0))"
+
+definition
+  CONS :: "['a item, 'a item] => 'a item" where
+  "CONS M N = In1(Scons M N)"
+
+inductive_set
+  list :: "'a item set => 'a item set"
+  for A :: "'a item set"
+  where
+    NIL_I:  "NIL: list A"
+  | CONS_I: "[| a: A;  M: list A |] ==> CONS a M : list A"
+
+
+typedef (List)
+    'a list = "list(range Leaf) :: 'a item set" 
+  by (blast intro: list.NIL_I)
+
+abbreviation "Case == Datatype.Case"
+abbreviation "Split == Datatype.Split"
+
+definition
+  List_case :: "['b, ['a item, 'a item]=>'b, 'a item] => 'b" where
+  "List_case c d = Case(%x. c)(Split(d))"
+  
+definition
+  List_rec  :: "['a item, 'b, ['a item, 'a item, 'b]=>'b] => 'b" where
+  "List_rec M c d = wfrec (pred_sexp^+)
+                           (%g. List_case c (%x y. d x y (g y))) M"
+
+
+(* *********************************************************************** *)
+(*                                                                         *)
+(* Abstracting data type                                                   *)
+(*                                                                         *)
+(* *********************************************************************** *)
+
+(*Declaring the abstract list constructors*)
+
+no_translations
+  "[x, xs]" == "x#[xs]"
+  "[x]" == "x#[]"
+no_notation
+  Nil  ("[]") and
+  Cons (infixr "#" 65)
+
+definition
+  Nil       :: "'a list"                               ("[]") where
+   "Nil  = Abs_List(NIL)"
+
+definition
+  "Cons"       :: "['a, 'a list] => 'a list"           (infixr "#" 65) where
+   "x#xs = Abs_List(CONS (Leaf x)(Rep_List xs))"
+
+definition
+  (* list Recursion -- the trancl is Essential; see list.ML *)
+  list_rec  :: "['a list, 'b, ['a, 'a list, 'b]=>'b] => 'b" where
+   "list_rec l c d =
+      List_rec(Rep_List l) c (%x y r. d(inv Leaf x)(Abs_List y) r)"
+
+definition
+  list_case :: "['b, ['a, 'a list]=>'b, 'a list] => 'b" where
+   "list_case a f xs = list_rec xs a (%x xs r. f x xs)"
+
+(* list Enumeration *)
+translations
+  "[x, xs]" == "x#[xs]"
+  "[x]"     == "x#[]"
+
+  "case xs of [] => a | y#ys => b" == "CONST list_case(a, %y ys. b, xs)"
+
+
+(* *********************************************************************** *)
+(*                                                                         *)
+(* Generalized Map Functionals                                             *)
+(*                                                                         *)
+(* *********************************************************************** *)
+
+  
+(* Generalized Map Functionals *)
+
+definition
+  Rep_map   :: "('b => 'a item) => ('b list => 'a item)" where
+   "Rep_map f xs = list_rec xs  NIL(%x l r. CONS(f x) r)"
+
+definition
+  Abs_map   :: "('a item => 'b) => 'a item => 'b list" where
+   "Abs_map g M  = List_rec M Nil (%N L r. g(N)#r)"
+
+
+definition
+  map       :: "('a=>'b) => ('a list => 'b list)" where
+  "map f xs = list_rec xs [] (%x l r. f(x)#r)"
+
+consts take      :: "['a list,nat] => 'a list"
+primrec
+  take_0:  "take xs 0 = []"
+  take_Suc: "take xs (Suc n) = list_case [] (%x l. x # take l n) xs"
+
+lemma ListI: "x : list (range Leaf) ==> x : List"
+by (simp add: List_def)
+
+lemma ListD: "x : List ==> x : list (range Leaf)"
+by (simp add: List_def)
+
+lemma list_unfold: "list(A) = usum {Numb(0)} (uprod A (list(A)))"
+  by (fast intro!: list.intros [unfolded NIL_def CONS_def]
+           elim: list.cases [unfolded NIL_def CONS_def])
+
+(*This justifies using list in other recursive type definitions*)
+lemma list_mono: "A<=B ==> list(A) <= list(B)"
+apply (rule subsetI)
+apply (erule list.induct)
+apply (auto intro!: list.intros)
+done
+
+(*Type checking -- list creates well-founded sets*)
+lemma list_sexp: "list(sexp) <= sexp"
+apply (rule subsetI)
+apply (erule list.induct)
+apply (unfold NIL_def CONS_def)
+apply (auto intro: sexp.intros sexp_In0I sexp_In1I)
+done
+
+(* A <= sexp ==> list(A) <= sexp *)
+lemmas list_subset_sexp = subset_trans [OF list_mono list_sexp] 
+
+
+(*Induction for the type 'a list *)
+lemma list_induct:
+    "[| P(Nil);    
+        !!x xs. P(xs) ==> P(x # xs) |]  ==> P(l)"
+apply (unfold Nil_def Cons_def) 
+apply (rule Rep_List_inverse [THEN subst])
+(*types force good instantiation*)
+apply (rule Rep_List [unfolded List_def, THEN list.induct], simp)
+apply (erule Abs_List_inverse [unfolded List_def, THEN subst], blast) 
+done
+
+
+(*** Isomorphisms ***)
+
+lemma inj_on_Abs_list: "inj_on Abs_List (list(range Leaf))"
+apply (rule inj_on_inverseI)
+apply (erule Abs_List_inverse [unfolded List_def])
+done
+
+(** Distinctness of constructors **)
+
+lemma CONS_not_NIL [iff]: "CONS M N ~= NIL"
+by (simp add: NIL_def CONS_def)
+
+lemmas NIL_not_CONS [iff] = CONS_not_NIL [THEN not_sym]
+lemmas CONS_neq_NIL = CONS_not_NIL [THEN notE, standard]
+lemmas NIL_neq_CONS = sym [THEN CONS_neq_NIL]
+
+lemma Cons_not_Nil [iff]: "x # xs ~= Nil"
+apply (unfold Nil_def Cons_def)
+apply (rule CONS_not_NIL [THEN inj_on_Abs_list [THEN inj_on_contraD]])
+apply (simp_all add: list.intros rangeI Rep_List [unfolded List_def])
+done
+
+lemmas Nil_not_Cons [iff] = Cons_not_Nil [THEN not_sym, standard]
+lemmas Cons_neq_Nil = Cons_not_Nil [THEN notE, standard]
+lemmas Nil_neq_Cons = sym [THEN Cons_neq_Nil]
+
+(** Injectiveness of CONS and Cons **)
+
+lemma CONS_CONS_eq [iff]: "(CONS K M)=(CONS L N) = (K=L & M=N)"
+by (simp add: CONS_def)
+
+(*For reasoning about abstract list constructors*)
+declare Rep_List [THEN ListD, intro] ListI [intro]
+declare list.intros [intro,simp]
+declare Leaf_inject [dest!]
+
+lemma Cons_Cons_eq [iff]: "(x#xs=y#ys) = (x=y & xs=ys)"
+apply (simp add: Cons_def)
+apply (subst Abs_List_inject)
+apply (auto simp add: Rep_List_inject)
+done
+
+lemmas Cons_inject2 = Cons_Cons_eq [THEN iffD1, THEN conjE, standard]
+
+lemma CONS_D: "CONS M N: list(A) ==> M: A & N: list(A)"
+  by (induct L == "CONS M N" set: list) auto
+
+lemma sexp_CONS_D: "CONS M N: sexp ==> M: sexp & N: sexp"
+apply (simp add: CONS_def In1_def)
+apply (fast dest!: Scons_D)
+done
+
+
+(*Reasoning about constructors and their freeness*)
+
+
+lemma not_CONS_self: "N: list(A) ==> !M. N ~= CONS M N"
+apply (erule list.induct) apply simp_all done
+
+lemma not_Cons_self2: "\<forall>x. l ~= x#l"
+by (induct l rule: list_induct) simp_all
+
+
+lemma neq_Nil_conv2: "(xs ~= []) = (\<exists>y ys. xs = y#ys)"
+by (induct xs rule: list_induct) auto
+
+(** Conversion rules for List_case: case analysis operator **)
+
+lemma List_case_NIL [simp]: "List_case c h NIL = c"
+by (simp add: List_case_def NIL_def)
+
+lemma List_case_CONS [simp]: "List_case c h (CONS M N) = h M N"
+by (simp add: List_case_def CONS_def)
+
+
+(*** List_rec -- by wf recursion on pred_sexp ***)
+
+(* The trancl(pred_sexp) is essential because pred_sexp_CONS_I1,2 would not
+   hold if pred_sexp^+ were changed to pred_sexp. *)
+
+lemma List_rec_unfold_lemma:
+     "(%M. List_rec M c d) == 
+      wfrec (pred_sexp^+) (%g. List_case c (%x y. d x y (g y)))"
+by (simp add: List_rec_def)
+
+lemmas List_rec_unfold = 
+    def_wfrec [OF List_rec_unfold_lemma wf_pred_sexp [THEN wf_trancl], 
+               standard]
+
+
+(** pred_sexp lemmas **)
+
+lemma pred_sexp_CONS_I1: 
+    "[| M: sexp;  N: sexp |] ==> (M, CONS M N) : pred_sexp^+"
+by (simp add: CONS_def In1_def)
+
+lemma pred_sexp_CONS_I2: 
+    "[| M: sexp;  N: sexp |] ==> (N, CONS M N) : pred_sexp^+"
+by (simp add: CONS_def In1_def)
+
+lemma pred_sexp_CONS_D: 
+    "(CONS M1 M2, N) : pred_sexp^+ ==>  
+     (M1,N) : pred_sexp^+ & (M2,N) : pred_sexp^+"
+apply (frule pred_sexp_subset_Sigma [THEN trancl_subset_Sigma, THEN subsetD])
+apply (blast dest!: sexp_CONS_D intro: pred_sexp_CONS_I1 pred_sexp_CONS_I2 
+                    trans_trancl [THEN transD])
+done
+
+
+(** Conversion rules for List_rec **)
+
+lemma List_rec_NIL [simp]: "List_rec NIL c h = c"
+apply (rule List_rec_unfold [THEN trans])
+apply (simp add: List_case_NIL)
+done
+
+lemma List_rec_CONS [simp]:
+     "[| M: sexp;  N: sexp |] 
+      ==> List_rec (CONS M N) c h = h M N (List_rec N c h)"
+apply (rule List_rec_unfold [THEN trans])
+apply (simp add: pred_sexp_CONS_I2)
+done
+
+
+(*** list_rec -- by List_rec ***)
+
+lemmas Rep_List_in_sexp =
+    subsetD [OF range_Leaf_subset_sexp [THEN list_subset_sexp]
+                Rep_List [THEN ListD]] 
+
+
+lemma list_rec_Nil [simp]: "list_rec Nil c h = c"
+by (simp add: list_rec_def ListI [THEN Abs_List_inverse] Nil_def)
+
+
+lemma list_rec_Cons [simp]: "list_rec (a#l) c h = h a l (list_rec l c h)"
+by (simp add: list_rec_def ListI [THEN Abs_List_inverse] Cons_def
+              Rep_List_inverse Rep_List [THEN ListD] inj_Leaf Rep_List_in_sexp)
+
+
+(*Type checking.  Useful?*)
+lemma List_rec_type:
+     "[| M: list(A);      
+         A<=sexp;         
+         c: C(NIL);       
+         !!x y r. [| x: A;  y: list(A);  r: C(y) |] ==> h x y r: C(CONS x y)  
+      |] ==> List_rec M c h : C(M :: 'a item)"
+apply (erule list.induct, simp) 
+apply (insert list_subset_sexp) 
+apply (subst List_rec_CONS, blast+)
+done
+
+
+
+(** Generalized map functionals **)
+
+lemma Rep_map_Nil [simp]: "Rep_map f Nil = NIL"
+by (simp add: Rep_map_def)
+
+lemma Rep_map_Cons [simp]: 
+    "Rep_map f(x#xs) = CONS(f x)(Rep_map f xs)"
+by (simp add: Rep_map_def)
+
+lemma Rep_map_type: "(!!x. f(x): A) ==> Rep_map f xs: list(A)"
+apply (simp add: Rep_map_def)
+apply (rule list_induct, auto)
+done
+
+lemma Abs_map_NIL [simp]: "Abs_map g NIL = Nil"
+by (simp add: Abs_map_def)
+
+lemma Abs_map_CONS [simp]: 
+    "[| M: sexp;  N: sexp |] ==> Abs_map g (CONS M N) = g(M) # Abs_map g N"
+by (simp add: Abs_map_def)
+
+(*Eases the use of primitive recursion.*)
+lemma def_list_rec_NilCons:
+    "[| !!xs. f(xs) = list_rec xs c h  |] 
+     ==> f [] = c & f(x#xs) = h x xs (f xs)"
+by simp 
+
+lemma Abs_map_inverse:
+     "[| M: list(A);  A<=sexp;  !!z. z: A ==> f(g(z)) = z |]  
+      ==> Rep_map f (Abs_map g M) = M"
+apply (erule list.induct, simp_all)
+apply (insert list_subset_sexp) 
+apply (subst Abs_map_CONS, blast)
+apply blast 
+apply simp 
+done
+
+(*Rep_map_inverse is obtained via Abs_Rep_map and map_ident*)
+
+(** list_case **)
+
+(* setting up rewrite sets *)
+
+text{*Better to have a single theorem with a conjunctive conclusion.*}
+declare def_list_rec_NilCons [OF list_case_def, simp]
+
+(** list_case **)
+
+lemma expand_list_case: 
+ "P(list_case a f xs) = ((xs=[] --> P a ) & (!y ys. xs=y#ys --> P(f y ys)))"
+by (induct xs rule: list_induct) simp_all
+
+
+(**** Function definitions ****)
+
+declare def_list_rec_NilCons [OF map_def, simp]
+
+(** The functional "map" and the generalized functionals **)
+
+lemma Abs_Rep_map: 
+     "(!!x. f(x): sexp) ==>  
+        Abs_map g (Rep_map f xs) = map (%t. g(f(t))) xs"
+apply (induct xs rule: list_induct)
+apply (simp_all add: Rep_map_type list_sexp [THEN subsetD])
+done
+
+
+(** Additional mapping lemmas **)
+
+lemma map_ident [simp]: "map(%x. x)(xs) = xs"
+by (induct xs rule: list_induct) simp_all
+
+lemma map_compose: "map(f o g)(xs) = map f (map g xs)"
+apply (simp add: o_def)
+apply (induct xs rule: list_induct)
+apply simp_all
+done
+
+
+(** take **)
+
+lemma take_Suc1 [simp]: "take [] (Suc x) = []"
+by simp
+
+lemma take_Suc2 [simp]: "take(a#xs)(Suc x) = a#take xs x"
+by simp
+
+lemma take_Nil [simp]: "take [] n = []"
+by (induct n) simp_all
+
+lemma take_take_eq [simp]: "\<forall>n. take (take xs n) n = take xs n"
+apply (induct xs rule: list_induct)
+apply simp_all
+apply (rule allI)
+apply (induct_tac n)
+apply auto
+done
+
+end
--- a/src/HOL/IsaMakefile	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/IsaMakefile	Mon Feb 22 11:19:15 2010 -0800
@@ -455,8 +455,8 @@
 $(LOG)/HOL-Induct.gz: $(OUT)/HOL Induct/Com.thy Induct/Comb.thy		\
   Induct/Ordinals.thy Induct/PropLog.thy Induct/QuoNestedDataType.thy	\
   Induct/QuoDataType.thy Induct/ROOT.ML Induct/Sexp.thy			\
-  Induct/Sigma_Algebra.thy Induct/ABexp.thy Induct/Term.thy		\
-  Induct/Tree.thy Induct/document/root.tex
+  Induct/Sigma_Algebra.thy Induct/ABexp.thy Induct/SList.thy		\
+  Induct/Term.thy Induct/Tree.thy Induct/document/root.tex
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Induct
 
 
--- a/src/HOL/Library/LaTeXsugar.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Library/LaTeXsugar.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -64,10 +64,10 @@
 consts DUMMY :: 'a ("\<^raw:\_>")
 
 (* THEOREMS *)
+notation (Rule output)
+  "==>"  ("\<^raw:\mbox{}\inferrule{\mbox{>_\<^raw:}}>\<^raw:{\mbox{>_\<^raw:}}>")
+
 syntax (Rule output)
-  "==>" :: "prop \<Rightarrow> prop \<Rightarrow> prop"
-  ("\<^raw:\mbox{}\inferrule{\mbox{>_\<^raw:}}>\<^raw:{\mbox{>_\<^raw:}}>")
-
   "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
   ("\<^raw:\mbox{}\inferrule{>_\<^raw:}>\<^raw:{\mbox{>_\<^raw:}}>")
 
@@ -76,21 +76,20 @@
 
   "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
 
-syntax (Axiom output)
-  "Trueprop" :: "bool \<Rightarrow> prop"
-  ("\<^raw:\mbox{}\inferrule{\mbox{}}{\mbox{>_\<^raw:}}>")
+notation (Axiom output)
+  "Trueprop"  ("\<^raw:\mbox{}\inferrule{\mbox{}}{\mbox{>_\<^raw:}}>")
 
+notation (IfThen output)
+  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
 syntax (IfThen output)
-  "==>" :: "prop \<Rightarrow> prop \<Rightarrow> prop"
-  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
   "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
   ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
   "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}> /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
   "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
 
+notation (IfThenNoBox output)
+  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
 syntax (IfThenNoBox output)
-  "==>" :: "prop \<Rightarrow> prop \<Rightarrow> prop"
-  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
   "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
   ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
   "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
--- a/src/HOL/Library/Multiset.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Library/Multiset.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -176,64 +176,6 @@
   by (simp add: multiset_eq_conv_count_eq)
 
 
-subsubsection {* Intersection *}
-
-definition multiset_inter :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" (infixl "#\<inter>" 70) where
-  "multiset_inter A B = A - (A - B)"
-
-lemma multiset_inter_count:
-  "count (A #\<inter> B) x = min (count A x) (count B x)"
-  by (simp add: multiset_inter_def multiset_typedef)
-
-lemma multiset_inter_commute: "A #\<inter> B = B #\<inter> A"
-  by (rule multi_count_ext) (simp add: multiset_inter_count)
-
-lemma multiset_inter_assoc: "A #\<inter> (B #\<inter> C) = A #\<inter> B #\<inter> C"
-  by (rule multi_count_ext) (simp add: multiset_inter_count)
-
-lemma multiset_inter_left_commute: "A #\<inter> (B #\<inter> C) = B #\<inter> (A #\<inter> C)"
-  by (rule multi_count_ext) (simp add: multiset_inter_count)
-
-lemmas multiset_inter_ac =
-  multiset_inter_commute
-  multiset_inter_assoc
-  multiset_inter_left_commute
-
-lemma multiset_inter_single: "a \<noteq> b \<Longrightarrow> {#a#} #\<inter> {#b#} = {#}"
-  by (rule multi_count_ext) (auto simp add: multiset_inter_count)
-
-lemma multiset_union_diff_commute:
-  assumes "B #\<inter> C = {#}"
-  shows "A + B - C = A - C + B"
-proof (rule multi_count_ext)
-  fix x
-  from assms have "min (count B x) (count C x) = 0"
-    by (auto simp add: multiset_inter_count multiset_eq_conv_count_eq)
-  then have "count B x = 0 \<or> count C x = 0"
-    by auto
-  then show "count (A + B - C) x = count (A - C + B) x"
-    by auto
-qed
-
-
-subsubsection {* Comprehension (filter) *}
-
-lemma count_MCollect [simp]:
-  "count {# x:#M. P x #} a = (if P a then count M a else 0)"
-  by (simp add: MCollect_def in_multiset multiset_typedef)
-
-lemma MCollect_empty [simp]: "MCollect {#} P = {#}"
-  by (rule multi_count_ext) simp
-
-lemma MCollect_single [simp]:
-  "MCollect {#x#} P = (if P x then {#x#} else {#})"
-  by (rule multi_count_ext) simp
-
-lemma MCollect_union [simp]:
-  "MCollect (M + N) f = MCollect M f + MCollect N f"
-  by (rule multi_count_ext) simp
-
-
 subsubsection {* Equality of multisets *}
 
 lemma single_not_empty [simp]: "{#a#} \<noteq> {#} \<and> {#} \<noteq> {#a#}"
@@ -335,64 +277,61 @@
 
 subsubsection {* Pointwise ordering induced by count *}
 
-definition mset_le :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool"  (infix "\<le>#" 50) where
-  "A \<le># B \<longleftrightarrow> (\<forall>a. count A a \<le> count B a)"
+instantiation multiset :: (type) ordered_ab_semigroup_add_imp_le
+begin
+
+definition less_eq_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
+  mset_le_def: "A \<le> B \<longleftrightarrow> (\<forall>a. count A a \<le> count B a)"
 
-definition mset_less :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool"  (infix "<#" 50) where
-  "A <# B \<longleftrightarrow> A \<le># B \<and> A \<noteq> B"
+definition less_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
+  mset_less_def: "(A::'a multiset) < B \<longleftrightarrow> A \<le> B \<and> A \<noteq> B"
 
-notation mset_le  (infix "\<subseteq>#" 50)
-notation mset_less  (infix "\<subset>#" 50)
+instance proof
+qed (auto simp add: mset_le_def mset_less_def multiset_eq_conv_count_eq intro: order_trans antisym)
+
+end
 
 lemma mset_less_eqI:
-  "(\<And>x. count A x \<le> count B x) \<Longrightarrow> A \<subseteq># B"
+  "(\<And>x. count A x \<le> count B x) \<Longrightarrow> A \<le> B"
   by (simp add: mset_le_def)
 
-lemma mset_le_refl[simp]: "A \<le># A"
-unfolding mset_le_def by auto
-
-lemma mset_le_trans: "A \<le># B \<Longrightarrow> B \<le># C \<Longrightarrow> A \<le># C"
-unfolding mset_le_def by (fast intro: order_trans)
-
-lemma mset_le_antisym: "A \<le># B \<Longrightarrow> B \<le># A \<Longrightarrow> A = B"
-apply (unfold mset_le_def)
-apply (rule multiset_eq_conv_count_eq [THEN iffD2])
-apply (blast intro: order_antisym)
-done
-
-lemma mset_le_exists_conv: "(A \<le># B) = (\<exists>C. B = A + C)"
+lemma mset_le_exists_conv:
+  "(A::'a multiset) \<le> B \<longleftrightarrow> (\<exists>C. B = A + C)"
 apply (unfold mset_le_def, rule iffI, rule_tac x = "B - A" in exI)
 apply (auto intro: multiset_eq_conv_count_eq [THEN iffD2])
 done
 
-lemma mset_le_mono_add_right_cancel[simp]: "(A + C \<le># B + C) = (A \<le># B)"
-unfolding mset_le_def by auto
+lemma mset_le_mono_add_right_cancel [simp]:
+  "(A::'a multiset) + C \<le> B + C \<longleftrightarrow> A \<le> B"
+  by (fact add_le_cancel_right)
 
-lemma mset_le_mono_add_left_cancel[simp]: "(C + A \<le># C + B) = (A \<le># B)"
-unfolding mset_le_def by auto
+lemma mset_le_mono_add_left_cancel [simp]:
+  "C + (A::'a multiset) \<le> C + B \<longleftrightarrow> A \<le> B"
+  by (fact add_le_cancel_left)
+
+lemma mset_le_mono_add:
+  "(A::'a multiset) \<le> B \<Longrightarrow> C \<le> D \<Longrightarrow> A + C \<le> B + D"
+  by (fact add_mono)
 
-lemma mset_le_mono_add: "\<lbrakk> A \<le># B; C \<le># D \<rbrakk> \<Longrightarrow> A + C \<le># B + D"
-apply (unfold mset_le_def)
-apply auto
-apply (erule_tac x = a in allE)+
-apply auto
-done
+lemma mset_le_add_left [simp]:
+  "(A::'a multiset) \<le> A + B"
+  unfolding mset_le_def by auto
+
+lemma mset_le_add_right [simp]:
+  "B \<le> (A::'a multiset) + B"
+  unfolding mset_le_def by auto
 
-lemma mset_le_add_left[simp]: "A \<le># A + B"
-unfolding mset_le_def by auto
-
-lemma mset_le_add_right[simp]: "B \<le># A + B"
-unfolding mset_le_def by auto
+lemma mset_le_single:
+  "a :# B \<Longrightarrow> {#a#} \<le> B"
+  by (simp add: mset_le_def)
 
-lemma mset_le_single: "a :# B \<Longrightarrow> {#a#} \<le># B"
-by (simp add: mset_le_def)
-
-lemma multiset_diff_union_assoc: "C \<le># B \<Longrightarrow> A + B - C = A + (B - C)"
-by (simp add: multiset_eq_conv_count_eq mset_le_def)
+lemma multiset_diff_union_assoc:
+  "C \<le> B \<Longrightarrow> (A::'a multiset) + B - C = A + (B - C)"
+  by (simp add: multiset_eq_conv_count_eq mset_le_def)
 
 lemma mset_le_multiset_union_diff_commute:
-assumes "B \<le># A"
-shows "A - B + C = A + C - B"
+  assumes "B \<le> A"
+  shows "(A::'a multiset) - B + C = A + C - B"
 proof -
   from mset_le_exists_conv [of "B" "A"] assms have "\<exists>D. A = B + D" ..
   from this obtain D where "A = B + D" ..
@@ -410,31 +349,19 @@
     done
 qed
 
-interpretation mset_order: order "op \<le>#" "op <#"
-proof qed (auto intro: order.intro mset_le_refl mset_le_antisym
-  mset_le_trans simp: mset_less_def)
-
-interpretation mset_order_cancel_semigroup:
-  ordered_cancel_ab_semigroup_add "op +" "op \<le>#" "op <#"
-proof qed (erule mset_le_mono_add [OF mset_le_refl])
-
-interpretation mset_order_semigroup_cancel:
-  ordered_ab_semigroup_add_imp_le "op +" "op \<le>#" "op <#"
-proof qed simp
-
-lemma mset_lessD: "A \<subset># B \<Longrightarrow> x \<in># A \<Longrightarrow> x \<in># B"
+lemma mset_lessD: "A < B \<Longrightarrow> x \<in># A \<Longrightarrow> x \<in># B"
 apply (clarsimp simp: mset_le_def mset_less_def)
 apply (erule_tac x=x in allE)
 apply auto
 done
 
-lemma mset_leD: "A \<subseteq># B \<Longrightarrow> x \<in># A \<Longrightarrow> x \<in># B"
+lemma mset_leD: "A \<le> B \<Longrightarrow> x \<in># A \<Longrightarrow> x \<in># B"
 apply (clarsimp simp: mset_le_def mset_less_def)
 apply (erule_tac x = x in allE)
 apply auto
 done
   
-lemma mset_less_insertD: "(A + {#x#} \<subset># B) \<Longrightarrow> (x \<in># B \<and> A \<subset># B)"
+lemma mset_less_insertD: "(A + {#x#} < B) \<Longrightarrow> (x \<in># B \<and> A < B)"
 apply (rule conjI)
  apply (simp add: mset_lessD)
 apply (clarsimp simp: mset_le_def mset_less_def)
@@ -443,30 +370,90 @@
  apply (auto split: split_if_asm)
 done
 
-lemma mset_le_insertD: "(A + {#x#} \<subseteq># B) \<Longrightarrow> (x \<in># B \<and> A \<subseteq># B)"
+lemma mset_le_insertD: "(A + {#x#} \<le> B) \<Longrightarrow> (x \<in># B \<and> A \<le> B)"
 apply (rule conjI)
  apply (simp add: mset_leD)
 apply (force simp: mset_le_def mset_less_def split: split_if_asm)
 done
 
-lemma mset_less_of_empty[simp]: "A \<subset># {#} \<longleftrightarrow> False"
+lemma mset_less_of_empty[simp]: "A < {#} \<longleftrightarrow> False"
   by (auto simp add: mset_less_def mset_le_def multiset_eq_conv_count_eq)
 
-lemma multi_psub_of_add_self[simp]: "A \<subset># A + {#x#}"
-by (auto simp: mset_le_def mset_less_def)
+lemma multi_psub_of_add_self[simp]: "A < A + {#x#}"
+  by (auto simp: mset_le_def mset_less_def)
 
-lemma multi_psub_self[simp]: "A \<subset># A = False"
-by (auto simp: mset_le_def mset_less_def)
+lemma multi_psub_self[simp]: "(A::'a multiset) < A = False"
+  by simp
 
 lemma mset_less_add_bothsides:
-  "T + {#x#} \<subset># S + {#x#} \<Longrightarrow> T \<subset># S"
-by (auto simp: mset_le_def mset_less_def)
+  "T + {#x#} < S + {#x#} \<Longrightarrow> T < S"
+  by (fact add_less_imp_less_right)
+
+lemma mset_less_empty_nonempty:
+  "{#} < S \<longleftrightarrow> S \<noteq> {#}"
+  by (auto simp: mset_le_def mset_less_def)
+
+lemma mset_less_diff_self:
+  "c \<in># B \<Longrightarrow> B - {#c#} < B"
+  by (auto simp: mset_le_def mset_less_def multiset_eq_conv_count_eq)
+
+
+subsubsection {* Intersection *}
+
+instantiation multiset :: (type) semilattice_inf
+begin
+
+definition inf_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" where
+  multiset_inter_def: "inf_multiset A B = A - (A - B)"
+
+instance proof -
+  have aux: "\<And>m n q :: nat. m \<le> n \<Longrightarrow> m \<le> q \<Longrightarrow> m \<le> n - (n - q)" by arith
+  show "OFCLASS('a multiset, semilattice_inf_class)" proof
+  qed (auto simp add: multiset_inter_def mset_le_def aux)
+qed
+
+end
+
+abbreviation multiset_inter :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" (infixl "#\<inter>" 70) where
+  "multiset_inter \<equiv> inf"
 
-lemma mset_less_empty_nonempty: "({#} \<subset># S) = (S \<noteq> {#})"
-by (auto simp: mset_le_def mset_less_def)
+lemma multiset_inter_count:
+  "count (A #\<inter> B) x = min (count A x) (count B x)"
+  by (simp add: multiset_inter_def multiset_typedef)
+
+lemma multiset_inter_single: "a \<noteq> b \<Longrightarrow> {#a#} #\<inter> {#b#} = {#}"
+  by (rule multi_count_ext) (auto simp add: multiset_inter_count)
 
-lemma mset_less_diff_self: "c \<in># B \<Longrightarrow> B - {#c#} \<subset># B"
-  by (auto simp: mset_le_def mset_less_def multiset_eq_conv_count_eq)
+lemma multiset_union_diff_commute:
+  assumes "B #\<inter> C = {#}"
+  shows "A + B - C = A - C + B"
+proof (rule multi_count_ext)
+  fix x
+  from assms have "min (count B x) (count C x) = 0"
+    by (auto simp add: multiset_inter_count multiset_eq_conv_count_eq)
+  then have "count B x = 0 \<or> count C x = 0"
+    by auto
+  then show "count (A + B - C) x = count (A - C + B) x"
+    by auto
+qed
+
+
+subsubsection {* Comprehension (filter) *}
+
+lemma count_MCollect [simp]:
+  "count {# x:#M. P x #} a = (if P a then count M a else 0)"
+  by (simp add: MCollect_def in_multiset multiset_typedef)
+
+lemma MCollect_empty [simp]: "MCollect {#} P = {#}"
+  by (rule multi_count_ext) simp
+
+lemma MCollect_single [simp]:
+  "MCollect {#x#} P = (if P x then {#x#} else {#})"
+  by (rule multi_count_ext) simp
+
+lemma MCollect_union [simp]:
+  "MCollect (M + N) f = MCollect M f + MCollect N f"
+  by (rule multi_count_ext) simp
 
 
 subsubsection {* Set of elements *}
@@ -657,7 +644,7 @@
 apply auto
 done
 
-lemma mset_less_size: "A \<subset># B \<Longrightarrow> size A < size B"
+lemma mset_less_size: "(A::'a multiset) < B \<Longrightarrow> size A < size B"
 proof (induct A arbitrary: B)
   case (empty M)
   then have "M \<noteq> {#}" by (simp add: mset_less_empty_nonempty)
@@ -666,12 +653,12 @@
   then show ?case by simp
 next
   case (add S x T)
-  have IH: "\<And>B. S \<subset># B \<Longrightarrow> size S < size B" by fact
-  have SxsubT: "S + {#x#} \<subset># T" by fact
-  then have "x \<in># T" and "S \<subset># T" by (auto dest: mset_less_insertD)
+  have IH: "\<And>B. S < B \<Longrightarrow> size S < size B" by fact
+  have SxsubT: "S + {#x#} < T" by fact
+  then have "x \<in># T" and "S < T" by (auto dest: mset_less_insertD)
   then obtain T' where T: "T = T' + {#x#}" 
     by (blast dest: multi_member_split)
-  then have "S \<subset># T'" using SxsubT 
+  then have "S < T'" using SxsubT 
     by (blast intro: mset_less_add_bothsides)
   then have "size S < size T'" using IH by simp
   then show ?case using T by simp
@@ -686,7 +673,7 @@
 
 definition
   mset_less_rel :: "('a multiset * 'a multiset) set" where
-  "mset_less_rel = {(A,B). A \<subset># B}"
+  "mset_less_rel = {(A,B). A < B}"
 
 lemma multiset_add_sub_el_shuffle: 
   assumes "c \<in># B" and "b \<noteq> c" 
@@ -709,29 +696,29 @@
 text {* The induction rules: *}
 
 lemma full_multiset_induct [case_names less]:
-assumes ih: "\<And>B. \<forall>A. A \<subset># B \<longrightarrow> P A \<Longrightarrow> P B"
+assumes ih: "\<And>B. \<forall>(A::'a multiset). A < B \<longrightarrow> P A \<Longrightarrow> P B"
 shows "P B"
 apply (rule wf_mset_less_rel [THEN wf_induct])
 apply (rule ih, auto simp: mset_less_rel_def)
 done
 
 lemma multi_subset_induct [consumes 2, case_names empty add]:
-assumes "F \<subseteq># A"
+assumes "F \<le> A"
   and empty: "P {#}"
   and insert: "\<And>a F. a \<in># A \<Longrightarrow> P F \<Longrightarrow> P (F + {#a#})"
 shows "P F"
 proof -
-  from `F \<subseteq># A`
+  from `F \<le> A`
   show ?thesis
   proof (induct F)
     show "P {#}" by fact
   next
     fix x F
-    assume P: "F \<subseteq># A \<Longrightarrow> P F" and i: "F + {#x#} \<subseteq># A"
+    assume P: "F \<le> A \<Longrightarrow> P F" and i: "F + {#x#} \<le> A"
     show "P (F + {#x#})"
     proof (rule insert)
       from i show "x \<in># A" by (auto dest: mset_le_insertD)
-      from i have "F \<subseteq># A" by (auto dest: mset_le_insertD)
+      from i have "F \<le> A" by (auto dest: mset_le_insertD)
       with P show "P F" .
     qed
   qed
@@ -875,12 +862,8 @@
     by (simp add: sorted_remove1 multiset_of_remove1 insort_remove1)
 qed
 
-lemma multiset_of_remdups_le: "multiset_of (remdups xs) \<le># multiset_of xs"
-apply (induct xs)
- apply auto
-apply (rule mset_le_trans)
- apply auto
-done
+lemma multiset_of_remdups_le: "multiset_of (remdups xs) \<le> multiset_of xs"
+  by (induct xs) (auto intro: order_trans)
 
 lemma multiset_of_update:
   "i < length ls \<Longrightarrow> multiset_of (ls[i := v]) = multiset_of ls - {#ls ! i#} + {#v#}"
@@ -969,7 +952,7 @@
   by (simp add: multiset_eq_conv_count_eq count_of_filter)
 
 lemma mset_less_eq_Bag [code]:
-  "Bag xs \<subseteq># A \<longleftrightarrow> (\<forall>(x, n) \<in> set xs. count_of xs x \<le> count A x)"
+  "Bag xs \<le> A \<longleftrightarrow> (\<forall>(x, n) \<in> set xs. count_of xs x \<le> count A x)"
     (is "?lhs \<longleftrightarrow> ?rhs")
 proof
   assume ?lhs then show ?rhs
@@ -990,10 +973,10 @@
 begin
 
 definition
-  "HOL.eq A B \<longleftrightarrow> A \<subseteq># B \<and> B \<subseteq># A"
+  "HOL.eq A B \<longleftrightarrow> (A::'a multiset) \<le> B \<and> B \<le> A"
 
 instance proof
-qed (simp add: eq_multiset_def mset_order.eq_iff)
+qed (simp add: eq_multiset_def eq_iff)
 
 end
 
@@ -1223,76 +1206,46 @@
 
 subsubsection {* Partial-order properties *}
 
-instantiation multiset :: (order) order
-begin
-
-definition less_multiset_def:
-  "M' < M \<longleftrightarrow> (M', M) \<in> mult {(x', x). x' < x}"
-
-definition le_multiset_def:
-  "M' <= M \<longleftrightarrow> M' = M \<or> M' < (M::'a multiset)"
-
-lemma trans_base_order: "trans {(x', x). x' < (x::'a::order)}"
-unfolding trans_def by (blast intro: order_less_trans)
-
-text {*
- \medskip Irreflexivity.
-*}
+definition less_multiset :: "'a\<Colon>order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "<#" 50) where
+  "M' <# M \<longleftrightarrow> (M', M) \<in> mult {(x', x). x' < x}"
 
-lemma mult_irrefl_aux:
-  "finite A ==> (\<forall>x \<in> A. \<exists>y \<in> A. x < (y::'a::order)) \<Longrightarrow> A = {}"
-by (induct rule: finite_induct) (auto intro: order_less_trans)
+definition le_multiset :: "'a\<Colon>order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "<=#" 50) where
+  "M' <=# M \<longleftrightarrow> M' <# M \<or> M' = M"
 
-lemma mult_less_not_refl: "\<not> M < (M::'a::order multiset)"
-apply (unfold less_multiset_def, auto)
-apply (drule trans_base_order [THEN mult_implies_one_step], auto)
-apply (drule finite_set_of [THEN mult_irrefl_aux [rule_format (no_asm)]])
-apply (simp add: set_of_eq_empty_iff)
-done
-
-lemma mult_less_irrefl [elim!]: "M < (M::'a::order multiset) ==> R"
-using insert mult_less_not_refl by fast
-
-
-text {* Transitivity. *}
+notation (xsymbol) less_multiset (infix "\<subset>#" 50)
+notation (xsymbol) le_multiset (infix "\<subseteq>#" 50)
 
-theorem mult_less_trans: "K < M ==> M < N ==> K < (N::'a::order multiset)"
-unfolding less_multiset_def mult_def by (blast intro: trancl_trans)
-
-text {* Asymmetry. *}
-
-theorem mult_less_not_sym: "M < N ==> \<not> N < (M::'a::order multiset)"
-apply auto
-apply (rule mult_less_not_refl [THEN notE])
-apply (erule mult_less_trans, assumption)
-done
-
-theorem mult_less_asym:
-  "M < N ==> (\<not> P ==> N < (M::'a::order multiset)) ==> P"
-using mult_less_not_sym by blast
-
-theorem mult_le_refl [iff]: "M <= (M::'a::order multiset)"
-unfolding le_multiset_def by auto
+interpretation multiset_order: order le_multiset less_multiset
+proof -
+  have irrefl: "\<And>M :: 'a multiset. \<not> M \<subset># M"
+  proof
+    fix M :: "'a multiset"
+    assume "M \<subset># M"
+    then have MM: "(M, M) \<in> mult {(x, y). x < y}" by (simp add: less_multiset_def)
+    have "trans {(x'::'a, x). x' < x}"
+      by (rule transI) simp
+    moreover note MM
+    ultimately have "\<exists>I J K. M = I + J \<and> M = I + K
+      \<and> J \<noteq> {#} \<and> (\<forall>k\<in>set_of K. \<exists>j\<in>set_of J. (k, j) \<in> {(x, y). x < y})"
+      by (rule mult_implies_one_step)
+    then obtain I J K where "M = I + J" and "M = I + K"
+      and "J \<noteq> {#}" and "(\<forall>k\<in>set_of K. \<exists>j\<in>set_of J. (k, j) \<in> {(x, y). x < y})" by blast
+    then have aux1: "K \<noteq> {#}" and aux2: "\<forall>k\<in>set_of K. \<exists>j\<in>set_of K. k < j" by auto
+    have "finite (set_of K)" by simp
+    moreover note aux2
+    ultimately have "set_of K = {}"
+      by (induct rule: finite_induct) (auto intro: order_less_trans)
+    with aux1 show False by simp
+  qed
+  have trans: "\<And>K M N :: 'a multiset. K \<subset># M \<Longrightarrow> M \<subset># N \<Longrightarrow> K \<subset># N"
+    unfolding less_multiset_def mult_def by (blast intro: trancl_trans)
+  show "order (le_multiset :: 'a multiset \<Rightarrow> _) less_multiset" proof
+  qed (auto simp add: le_multiset_def irrefl dest: trans)
+qed
 
-text {* Anti-symmetry. *}
-
-theorem mult_le_antisym:
-  "M <= N ==> N <= M ==> M = (N::'a::order multiset)"
-unfolding le_multiset_def by (blast dest: mult_less_not_sym)
-
-text {* Transitivity. *}
-
-theorem mult_le_trans:
-  "K <= M ==> M <= N ==> K <= (N::'a::order multiset)"
-unfolding le_multiset_def by (blast intro: mult_less_trans)
-
-theorem mult_less_le: "(M < N) = (M <= N \<and> M \<noteq> (N::'a::order multiset))"
-unfolding le_multiset_def by auto
-
-instance proof
-qed (auto simp add: mult_less_le dest: mult_le_antisym elim: mult_le_trans)
-
-end
+lemma mult_less_irrefl [elim!]:
+  "M \<subset># (M::'a::order multiset) ==> R"
+  by (simp add: multiset_order.less_irrefl)
 
 
 subsubsection {* Monotonicity of multiset union *}
@@ -1306,52 +1259,26 @@
 apply (simp add: add_assoc)
 done
 
-lemma union_less_mono2: "B < D ==> C + B < C + (D::'a::order multiset)"
+lemma union_less_mono2: "B \<subset># D ==> C + B \<subset># C + (D::'a::order multiset)"
 apply (unfold less_multiset_def mult_def)
 apply (erule trancl_induct)
  apply (blast intro: mult1_union transI order_less_trans r_into_trancl)
 apply (blast intro: mult1_union transI order_less_trans r_into_trancl trancl_trans)
 done
 
-lemma union_less_mono1: "B < D ==> B + C < D + (C::'a::order multiset)"
+lemma union_less_mono1: "B \<subset># D ==> B + C \<subset># D + (C::'a::order multiset)"
 apply (subst add_commute [of B C])
 apply (subst add_commute [of D C])
 apply (erule union_less_mono2)
 done
 
 lemma union_less_mono:
-  "A < C ==> B < D ==> A + B < C + (D::'a::order multiset)"
-by (blast intro!: union_less_mono1 union_less_mono2 mult_less_trans)
-
-lemma union_le_mono:
-  "A <= C ==> B <= D ==> A + B <= C + (D::'a::order multiset)"
-unfolding le_multiset_def
-by (blast intro: union_less_mono union_less_mono1 union_less_mono2)
+  "A \<subset># C ==> B \<subset># D ==> A + B \<subset># C + (D::'a::order multiset)"
+  by (blast intro!: union_less_mono1 union_less_mono2 multiset_order.less_trans)
 
-lemma empty_leI [iff]: "{#} <= (M::'a::order multiset)"
-apply (unfold le_multiset_def less_multiset_def)
-apply (case_tac "M = {#}")
- prefer 2
- apply (subgoal_tac "({#} + {#}, {#} + M) \<in> mult (Collect (split op <))")
-  prefer 2
-  apply (rule one_step_implies_mult)
-    apply (simp only: trans_def)
-    apply auto
-done
-
-lemma union_upper1: "A <= A + (B::'a::order multiset)"
-proof -
-  have "A + {#} <= A + B" by (blast intro: union_le_mono)
-  then show ?thesis by simp
-qed
-
-lemma union_upper2: "B <= A + (B::'a::order multiset)"
-by (subst add_commute) (rule union_upper1)
-
-instance multiset :: (order) ordered_ab_semigroup_add
-apply intro_classes
-apply (erule union_le_mono[OF mult_le_refl])
-done
+interpretation multiset_order: ordered_ab_semigroup_add plus le_multiset less_multiset
+proof
+qed (auto simp add: le_multiset_def intro: union_less_mono2)
 
 
 subsection {* The fold combinator *}
@@ -1406,7 +1333,7 @@
   "fold_msetG f z A x \<Longrightarrow> fold_msetG f z A y \<Longrightarrow> y = x"
 proof (induct arbitrary: x y z rule: full_multiset_induct)
   case (less M x\<^isub>1 x\<^isub>2 Z)
-  have IH: "\<forall>A. A \<subset># M \<longrightarrow> 
+  have IH: "\<forall>A. A < M \<longrightarrow> 
     (\<forall>x x' x''. fold_msetG f x'' A x \<longrightarrow> fold_msetG f x'' A x'
                \<longrightarrow> x' = x)" by fact
   have Mfoldx\<^isub>1: "fold_msetG f Z M x\<^isub>1" and Mfoldx\<^isub>2: "fold_msetG f Z M x\<^isub>2" by fact+
@@ -1426,8 +1353,8 @@
       fix C c v
       assume "M = C + {#c#}" and "x\<^isub>2 = f c v" and Cv: "fold_msetG f Z C v"
       then have MCc: "M = C + {#c#}" and x\<^isub>2: "x\<^isub>2 = f c v" by auto
-      then have CsubM: "C \<subset># M" by simp
-      from MBb have BsubM: "B \<subset># M" by simp
+      then have CsubM: "C < M" by simp
+      from MBb have BsubM: "B < M" by simp
       show ?case
       proof cases
         assume "b=c"
@@ -1438,8 +1365,8 @@
         let ?D = "B - {#c#}"
         have cinB: "c \<in># B" and binC: "b \<in># C" using MBb MCc diff
           by (auto intro: insert_noteq_member dest: sym)
-        have "B - {#c#} \<subset># B" using cinB by (rule mset_less_diff_self)
-        then have DsubM: "?D \<subset># M" using BsubM by (blast intro: mset_order.less_trans)
+        have "B - {#c#} < B" using cinB by (rule mset_less_diff_self)
+        then have DsubM: "?D < M" using BsubM by (blast intro: order_less_trans)
         from MBb MCc have "B + {#b#} = C + {#c#}" by blast
         then have [simp]: "B + {#b#} - {#c#} = C"
           using MBb MCc binC cinB by auto
@@ -1769,6 +1696,37 @@
 lemma multi_union_self_other_eq: "(A::'a multiset) + X = A + Y \<Longrightarrow> X = Y"
   by (fact add_imp_eq)
 
-lemmas mset_less_trans = mset_order.less_trans
+lemma mset_less_trans: "(M::'a multiset) < K \<Longrightarrow> K < N \<Longrightarrow> M < N"
+  by (fact order_less_trans)
+
+lemma multiset_inter_commute: "A #\<inter> B = B #\<inter> A"
+  by (fact inf.commute)
+
+lemma multiset_inter_assoc: "A #\<inter> (B #\<inter> C) = A #\<inter> B #\<inter> C"
+  by (fact inf.assoc [symmetric])
+
+lemma multiset_inter_left_commute: "A #\<inter> (B #\<inter> C) = B #\<inter> (A #\<inter> C)"
+  by (fact inf.left_commute)
+
+lemmas multiset_inter_ac =
+  multiset_inter_commute
+  multiset_inter_assoc
+  multiset_inter_left_commute
+
+lemma mult_less_not_refl:
+  "\<not> M \<subset># (M::'a::order multiset)"
+  by (fact multiset_order.less_irrefl)
+
+lemma mult_less_trans:
+  "K \<subset># M ==> M \<subset># N ==> K \<subset># (N::'a::order multiset)"
+  by (fact multiset_order.less_trans)
+    
+lemma mult_less_not_sym:
+  "M \<subset># N ==> \<not> N \<subset># (M::'a::order multiset)"
+  by (fact multiset_order.less_not_sym)
+
+lemma mult_less_asym:
+  "M \<subset># N ==> (\<not> P ==> N \<subset># (M::'a::order multiset)) ==> P"
+  by (fact multiset_order.less_asym)
 
 end
\ No newline at end of file
--- a/src/HOL/Library/OptionalSugar.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Library/OptionalSugar.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -37,15 +37,15 @@
 
 (* Let *)
 translations 
-  "_bind (p,DUMMY) e" <= "_bind p (CONST fst e)"
-  "_bind (DUMMY,p) e" <= "_bind p (CONST snd e)"
+  "_bind (p, CONST DUMMY) e" <= "_bind p (CONST fst e)"
+  "_bind (CONST DUMMY, p) e" <= "_bind p (CONST snd e)"
 
   "_tuple_args x (_tuple_args y z)" ==
     "_tuple_args x (_tuple_arg (_tuple y z))"
 
-  "_bind (Some p) e" <= "_bind p (CONST the e)"
-  "_bind (p#DUMMY) e" <= "_bind p (CONST hd e)"
-  "_bind (DUMMY#p) e" <= "_bind p (CONST tl e)"
+  "_bind (CONST Some p) e" <= "_bind p (CONST the e)"
+  "_bind (p # CONST DUMMY) e" <= "_bind p (CONST hd e)"
+  "_bind (CONST DUMMY # p) e" <= "_bind p (CONST tl e)"
 
 (* type constraints with spacing *)
 setup {*
--- a/src/HOL/Library/Permutation.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Library/Permutation.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -154,7 +154,7 @@
   done
 
 lemma multiset_of_le_perm_append:
-    "(multiset_of xs \<le># multiset_of ys) = (\<exists>zs. xs @ zs <~~> ys)"
+    "multiset_of xs \<le> multiset_of ys \<longleftrightarrow> (\<exists>zs. xs @ zs <~~> ys)"
   apply (auto simp: multiset_of_eq_perm[THEN sym] mset_le_exists_conv)
   apply (insert surj_multiset_of, drule surjD)
   apply (blast intro: sym)+
--- a/src/HOL/Library/Quotient_Sum.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Library/Quotient_Sum.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -24,7 +24,7 @@
 declare [[map "+" = (sum_map, sum_rel)]]
 
 
-text {* should probably be in Sum_Type.thy *}
+text {* should probably be in @{theory Sum_Type} *}
 lemma split_sum_all:
   shows "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (Inl x)) \<and> (\<forall>x. P (Inr x))"
   apply(auto)
--- a/src/HOL/Library/SetsAndFunctions.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Library/SetsAndFunctions.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -119,14 +119,14 @@
   apply (force simp add: mult_assoc)
   done
 
-interpretation set_comm_monoid_add: comm_monoid_add "{0}" "op \<oplus> :: ('a::comm_monoid_add) set => 'a set => 'a set"
+interpretation set_comm_monoid_add: comm_monoid_add "op \<oplus> :: ('a::comm_monoid_add) set => 'a set => 'a set" "{0}"
   apply default
    apply (unfold set_plus_def)
    apply (force simp add: add_ac)
   apply force
   done
 
-interpretation set_comm_monoid_mult: comm_monoid_mult "{1}" "op \<otimes> :: ('a::comm_monoid_mult) set => 'a set => 'a set"
+interpretation set_comm_monoid_mult: comm_monoid_mult "op \<otimes> :: ('a::comm_monoid_mult) set => 'a set => 'a set" "{1}"
   apply default
    apply (unfold set_times_def)
    apply (force simp add: mult_ac)
--- a/src/HOL/List.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/List.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -358,11 +358,11 @@
 
 parse_translation (advanced) {*
 let
-  val NilC = Syntax.const @{const_name Nil};
-  val ConsC = Syntax.const @{const_name Cons};
-  val mapC = Syntax.const @{const_name map};
-  val concatC = Syntax.const @{const_name concat};
-  val IfC = Syntax.const @{const_name If};
+  val NilC = Syntax.const @{const_syntax Nil};
+  val ConsC = Syntax.const @{const_syntax Cons};
+  val mapC = Syntax.const @{const_syntax map};
+  val concatC = Syntax.const @{const_syntax concat};
+  val IfC = Syntax.const @{const_syntax If};
 
   fun singl x = ConsC $ x $ NilC;
 
@@ -371,12 +371,14 @@
       val x = Free (Name.variant (fold Term.add_free_names [p, e] []) "x", dummyT);
       val e = if opti then singl e else e;
       val case1 = Syntax.const @{syntax_const "_case1"} $ p $ e;
-      val case2 = Syntax.const @{syntax_const "_case1"} $ Syntax.const Term.dummy_patternN $ NilC;
+      val case2 =
+        Syntax.const @{syntax_const "_case1"} $
+          Syntax.const @{const_syntax dummy_pattern} $ NilC;
       val cs = Syntax.const @{syntax_const "_case2"} $ case1 $ case2;
       val ft = Datatype_Case.case_tr false Datatype.info_of_constr ctxt [x, cs];
     in lambda x ft end;
 
-  fun abs_tr ctxt (p as Free(s,T)) e opti =
+  fun abs_tr ctxt (p as Free (s, T)) e opti =
         let
           val thy = ProofContext.theory_of ctxt;
           val s' = Sign.intern_const thy s;
@@ -1677,12 +1679,23 @@
 lemma hd_drop_conv_nth: "\<lbrakk> xs \<noteq> []; n < length xs \<rbrakk> \<Longrightarrow> hd(drop n xs) = xs!n"
 by(simp add: hd_conv_nth)
 
+lemma set_take_subset_set_take:
+  "m <= n \<Longrightarrow> set(take m xs) <= set(take n xs)"
+by(induct xs arbitrary: m n)(auto simp:take_Cons split:nat.split)
+
 lemma set_take_subset: "set(take n xs) \<subseteq> set xs"
 by(induct xs arbitrary: n)(auto simp:take_Cons split:nat.split)
 
 lemma set_drop_subset: "set(drop n xs) \<subseteq> set xs"
 by(induct xs arbitrary: n)(auto simp:drop_Cons split:nat.split)
 
+lemma set_drop_subset_set_drop:
+  "m >= n \<Longrightarrow> set(drop m xs) <= set(drop n xs)"
+apply(induct xs arbitrary: m n)
+apply(auto simp:drop_Cons split:nat.split)
+apply (metis set_drop_subset subset_iff)
+done
+
 lemma in_set_takeD: "x : set(take n xs) \<Longrightarrow> x : set xs"
 using set_take_subset by fast
 
@@ -2352,7 +2365,7 @@
 proof (induct xss arbitrary: xs)
   case Nil show ?case by simp
 next
-  interpret monoid_add "[]" "op @" proof qed simp_all
+  interpret monoid_add "op @" "[]" proof qed simp_all
   case Cons then show ?case by (simp add: foldl_absorb0)
 qed
 
--- a/src/HOL/Main.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Main.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -1,7 +1,7 @@
 header {* Main HOL *}
 
 theory Main
-imports Plain Predicate_Compile Nitpick Quotient
+imports Plain Predicate_Compile Nitpick
 begin
 
 text {*
--- a/src/HOL/MicroJava/DFA/Semilat.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/MicroJava/DFA/Semilat.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -22,16 +22,17 @@
   "lesssub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool"
   "plussub" :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c" 
 (*<*)
-syntax 
-  "lesub"   :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /<='__ _)" [50, 1000, 51] 50)
-  "lesssub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /<'__ _)" [50, 1000, 51] 50)
-  "plussub" :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c" ("(_ /+'__ _)" [65, 1000, 66] 65)
+notation
+  "lesub"  ("(_ /<='__ _)" [50, 1000, 51] 50) and
+  "lesssub"  ("(_ /<'__ _)" [50, 1000, 51] 50) and
+  "plussub"  ("(_ /+'__ _)" [65, 1000, 66] 65)
 (*>*)
-syntax (xsymbols)
-  "lesub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubseteq>\<^bsub>_\<^esub> _)" [50, 0, 51] 50)
-  "lesssub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubset>\<^bsub>_\<^esub> _)" [50, 0, 51] 50)
-  "plussub" :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c" ("(_ /\<squnion>\<^bsub>_\<^esub> _)" [65, 0, 66] 65)
+notation (xsymbols)
+  "lesub"  ("(_ /\<sqsubseteq>\<^bsub>_\<^esub> _)" [50, 0, 51] 50) and
+  "lesssub"  ("(_ /\<sqsubset>\<^bsub>_\<^esub> _)" [50, 0, 51] 50) and
+  "plussub"  ("(_ /\<squnion>\<^bsub>_\<^esub> _)" [65, 0, 66] 65)
 (*<*)
+syntax
  (* allow \<sub> instead of \<bsub>..\<esub> *)  
   "_lesub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubseteq>\<^sub>_ _)" [50, 1000, 51] 50)
   "_lesssub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubset>\<^sub>_ _)" [50, 1000, 51] 50)
--- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -1,4 +1,4 @@
-(*  Title:      Library/Euclidean_Space
+(*  Title:      Library/Multivariate_Analysis/Euclidean_Space.thy
     Author:     Amine Chaieb, University of Cambridge
 *)
 
@@ -66,8 +66,8 @@
 
 instantiation cart :: (plus,finite) plus
 begin
-definition  vector_add_def : "op + \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) + (y$i)))"
-instance ..
+  definition  vector_add_def : "op + \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) + (y$i)))"
+  instance ..
 end
 
 instantiation cart :: (times,finite) times
@@ -76,39 +76,42 @@
   instance ..
 end
 
-instantiation cart :: (minus,finite) minus begin
+instantiation cart :: (minus,finite) minus
+begin
   definition vector_minus_def : "op - \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) - (y$i)))"
-instance ..
+  instance ..
 end
 
-instantiation cart :: (uminus,finite) uminus begin
+instantiation cart :: (uminus,finite) uminus
+begin
   definition vector_uminus_def : "uminus \<equiv> (\<lambda> x.  (\<chi> i. - (x$i)))"
-instance ..
+  instance ..
 end
 
-instantiation cart :: (zero,finite) zero begin
+instantiation cart :: (zero,finite) zero
+begin
   definition vector_zero_def : "0 \<equiv> (\<chi> i. 0)"
-instance ..
+  instance ..
 end
 
-instantiation cart :: (one,finite) one begin
+instantiation cart :: (one,finite) one
+begin
   definition vector_one_def : "1 \<equiv> (\<chi> i. 1)"
-instance ..
+  instance ..
 end
 
 instantiation cart :: (ord,finite) ord
- begin
-definition vector_le_def:
-  "less_eq (x :: 'a ^'b) y = (ALL i. x$i <= y$i)"
-definition vector_less_def: "less (x :: 'a ^'b) y = (ALL i. x$i < y$i)"
-
-instance by (intro_classes)
+begin
+  definition vector_le_def:
+    "less_eq (x :: 'a ^'b) y = (ALL i. x$i <= y$i)"
+  definition vector_less_def: "less (x :: 'a ^'b) y = (ALL i. x$i < y$i)"
+  instance by (intro_classes)
 end
 
 instantiation cart :: (scaleR, finite) scaleR
 begin
-definition vector_scaleR_def: "scaleR = (\<lambda> r x.  (\<chi> i. scaleR r (x$i)))"
-instance ..
+  definition vector_scaleR_def: "scaleR = (\<lambda> r x.  (\<chi> i. scaleR r (x$i)))"
+  instance ..
 end
 
 text{* Also the scalar-vector multiplication. *}
--- a/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -1,11 +1,11 @@
-(* Title:      HOL/Library/Finite_Cartesian_Product
-   Author:     Amine Chaieb, University of Cambridge
+(*  Title:      HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy
+    Author:     Amine Chaieb, University of Cambridge
 *)
 
 header {* Definition of finite Cartesian product types. *}
 
 theory Finite_Cartesian_Product
-imports Main (*FIXME: ATP_Linkup is only needed for metis at a few places. We could dispense of that by changing the proofs.*)
+imports Main
 begin
 
 subsection {* Finite Cartesian products, with indexing and lambdas. *}
@@ -14,9 +14,9 @@
   ('a, 'b) cart = "UNIV :: (('b::finite) \<Rightarrow> 'a) set"
   morphisms Cart_nth Cart_lambda ..
 
-notation Cart_nth (infixl "$" 90)
-
-notation (xsymbols) Cart_lambda (binder "\<chi>" 10)
+notation
+  Cart_nth (infixl "$" 90) and
+  Cart_lambda (binder "\<chi>" 10)
 
 (*
   Translate "'b ^ 'n" into "'b ^ ('n :: finite)". When 'n has already more than
@@ -39,10 +39,7 @@
 *}
 
 lemma stupid_ext: "(\<forall>x. f x = g x) \<longleftrightarrow> (f = g)"
-  apply auto
-  apply (rule ext)
-  apply auto
-  done
+  by (auto intro: ext)
 
 lemma Cart_eq: "(x = y) \<longleftrightarrow> (\<forall>i. x$i = y$i)"
   by (simp add: Cart_nth_inject [symmetric] expand_fun_eq)
@@ -75,7 +72,10 @@
   unfolding sndcart_def by simp
 
 lemma finite_sum_image: "(UNIV::('a + 'b) set) = range Inl \<union> range Inr"
-by (auto, case_tac x, auto)
+  apply auto
+  apply (case_tac x)
+   apply auto
+  done
 
 lemma fstcart_pastecart[simp]: "fstcart (pastecart x y) = x"
   by (simp add: Cart_eq)
@@ -90,9 +90,9 @@
   using pastecart_fst_snd[of x] pastecart_fst_snd[of y] by metis
 
 lemma forall_pastecart: "(\<forall>p. P p) \<longleftrightarrow> (\<forall>x y. P (pastecart x y))"
-  by (metis pastecart_fst_snd fstcart_pastecart sndcart_pastecart)
+  by (metis pastecart_fst_snd)
 
-lemma exists_pastecart: "(\<exists>p. P p)  \<longleftrightarrow> (\<exists>x y. P (pastecart x y))"
-  by (metis pastecart_fst_snd fstcart_pastecart sndcart_pastecart)
+lemma exists_pastecart: "(\<exists>p. P p) \<longleftrightarrow> (\<exists>x y. P (pastecart x y))"
+  by (metis pastecart_fst_snd)
 
 end
--- a/src/HOL/Mutabelle/Mutabelle.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Mutabelle/Mutabelle.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -14,7 +14,7 @@
   (@{const_name HOL.undefined}, "'a"),
   (@{const_name HOL.default}, "'a"),
   (@{const_name dummy_pattern}, "'a::{}"),
-  (@{const_name Algebras.uminus}, "'a"),
+  (@{const_name Groups.uminus}, "'a"),
   (@{const_name Nat.size}, "'a"),
   (@{const_name Groups.abs}, "'a")];
 
--- a/src/HOL/NSA/NSA.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/NSA/NSA.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -671,12 +671,12 @@
   0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*)
 fun reorient_proc sg _ (_ $ t $ u) =
   case u of
-      Const(@{const_name Algebras.zero}, _) => NONE
-    | Const(@{const_name Algebras.one}, _) => NONE
+      Const(@{const_name Groups.zero}, _) => NONE
+    | Const(@{const_name Groups.one}, _) => NONE
     | Const(@{const_name Int.number_of}, _) $ _ => NONE
     | _ => SOME (case t of
-                Const(@{const_name Algebras.zero}, _) => meta_zero_approx_reorient
-              | Const(@{const_name Algebras.one}, _) => meta_one_approx_reorient
+                Const(@{const_name Groups.zero}, _) => meta_zero_approx_reorient
+              | Const(@{const_name Groups.one}, _) => meta_one_approx_reorient
               | Const(@{const_name Int.number_of}, _) $ _ =>
                                  meta_number_of_approx_reorient);
 
--- a/src/HOL/Nitpick.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Nitpick.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -8,7 +8,7 @@
 header {* Nitpick: Yet Another Counterexample Generator for Isabelle/HOL *}
 
 theory Nitpick
-imports Map SAT
+imports Map Quotient SAT
 uses ("Tools/Nitpick/kodkod.ML")
      ("Tools/Nitpick/kodkod_sat.ML")
      ("Tools/Nitpick/nitpick_util.ML")
--- a/src/HOL/Nitpick_Examples/Core_Nits.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Core_Nits.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -11,82 +11,67 @@
 imports Main
 begin
 
-nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60 s]
+nitpick_params [max_potential = 0, sat_solver = MiniSat_JNI, max_threads = 1,
+                timeout = 60 s]
 
 subsection {* Curry in a Hurry *}
 
 lemma "(\<lambda>f x y. (curry o split) f x y) = (\<lambda>f x y. (\<lambda>x. x) f x y)"
-nitpick [card = 1\<midarrow>4, expect = none]
-nitpick [card = 100, expect = none, timeout = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 by auto
 
 lemma "(\<lambda>f p. (split o curry) f p) = (\<lambda>f p. (\<lambda>x. x) f p)"
-nitpick [card = 2]
-nitpick [card = 1\<midarrow>4, expect = none]
-nitpick [card = 10, expect = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 by auto
 
 lemma "split (curry f) = f"
-nitpick [card = 1\<midarrow>4, expect = none]
-nitpick [card = 10, expect = none]
-nitpick [card = 40, expect = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 by auto
 
 lemma "curry (split f) = f"
-nitpick [card = 1\<midarrow>4, expect = none]
-nitpick [card = 40, expect = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 by auto
 
 lemma "(split o curry) f = f"
-nitpick [card = 1\<midarrow>4, expect = none]
-nitpick [card = 40, expect = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 by auto
 
 lemma "(curry o split) f = f"
-nitpick [card = 1\<midarrow>4, expect = none]
-nitpick [card = 1000, expect = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 by auto
 
 lemma "(split o curry) f = (\<lambda>x. x) f"
-nitpick [card = 1\<midarrow>4, expect = none]
-nitpick [card = 40, expect = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 by auto
 
 lemma "(curry o split) f = (\<lambda>x. x) f"
-nitpick [card = 1\<midarrow>4, expect = none]
-nitpick [card = 40, expect = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 by auto
 
 lemma "((split o curry) f) p = ((\<lambda>x. x) f) p"
-nitpick [card = 1\<midarrow>4, expect = none]
-nitpick [card = 40, expect = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 by auto
 
 lemma "((curry o split) f) x = ((\<lambda>x. x) f) x"
-nitpick [card = 1\<midarrow>4, expect = none]
-nitpick [card = 1000, expect = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 by auto
 
 lemma "((curry o split) f) x y = ((\<lambda>x. x) f) x y"
-nitpick [card = 1\<midarrow>4, expect = none]
-nitpick [card = 1000, expect = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 by auto
 
 lemma "split o curry = (\<lambda>x. x)"
-nitpick [card = 1\<midarrow>4, expect = none]
-nitpick [card = 40, expect = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 apply (rule ext)+
 by auto
 
 lemma "curry o split = (\<lambda>x. x)"
-nitpick [card = 1\<midarrow>4, expect = none]
-nitpick [card = 100, expect = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 apply (rule ext)+
 by auto
 
 lemma "split (\<lambda>x y. f (x, y)) = f"
-nitpick [card = 1\<midarrow>4, expect = none]
-nitpick [card = 40, expect = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 by auto
 
 subsection {* Representations *}
@@ -96,13 +81,12 @@
 by auto
 
 lemma "(\<exists>g. \<forall>x. g (f x) = x) \<longrightarrow> (\<forall>y. \<exists>x. y = f x)"
-nitpick [card 'a = 35, card 'b = 34, expect = genuine]
-nitpick [card = 1\<midarrow>15, mono, expect = none]
+nitpick [card 'a = 25, card 'b = 24, expect = genuine]
+nitpick [card = 1\<midarrow>10, mono, expect = none]
 oops
 
 lemma "\<exists>f. f = (\<lambda>x. x) \<and> f y \<noteq> y"
 nitpick [card = 1, expect = genuine]
-nitpick [card = 2, expect = genuine]
 nitpick [card = 5, expect = genuine]
 oops
 
@@ -112,8 +96,7 @@
 oops
 
 lemma "{(a\<Colon>'a\<times>'a, b\<Colon>'b)}^-1 = {(b, a)}"
-nitpick [card = 1\<midarrow>6, expect = none]
-nitpick [card = 20, expect = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 by auto
 
 lemma "fst (a, b) = a"
@@ -121,7 +104,7 @@
 by auto
 
 lemma "\<exists>P. P = Id"
-nitpick [card = 1\<midarrow>4, expect = none]
+nitpick [card = 1\<midarrow>20, expect = none]
 by auto
 
 lemma "(a\<Colon>'a\<Rightarrow>'b, a) \<in> Id\<^sup>*"
@@ -129,11 +112,11 @@
 by auto
 
 lemma "(a\<Colon>'a\<times>'a, a) \<in> Id\<^sup>* \<union> {(a, b)}\<^sup>*"
-nitpick [card = 1\<midarrow>6, expect = none]
+nitpick [card = 1\<midarrow>4, expect = none]
 by auto
 
 lemma "Id (a, a)"
-nitpick [card = 1\<midarrow>100, expect = none]
+nitpick [card = 1\<midarrow>50, expect = none]
 by (auto simp: Id_def Collect_def)
 
 lemma "Id ((a\<Colon>'a, b\<Colon>'a), (a, b))"
@@ -151,7 +134,7 @@
 lemma "g = Let (A \<or> B)"
 nitpick [card = 1, expect = none]
 nitpick [card = 2, expect = genuine]
-nitpick [card = 20, expect = genuine]
+nitpick [card = 12, expect = genuine]
 oops
 
 lemma "(let a_or_b = A \<or> B in a_or_b \<or> \<not> a_or_b)"
@@ -172,37 +155,30 @@
 
 lemma "(a\<Colon>'a\<times>'a, a\<Colon>'a\<times>'a) \<in> R"
 nitpick [card = 1, expect = genuine]
-nitpick [card = 2, expect = genuine]
-nitpick [card = 4, expect = genuine]
 nitpick [card = 20, expect = genuine]
-nitpick [card = 10, dont_box, expect = genuine]
+nitpick [card = 5, dont_box, expect = genuine]
 oops
 
 lemma "f (g\<Colon>'a\<Rightarrow>'a) = x"
 nitpick [card = 3, expect = genuine]
 nitpick [card = 3, dont_box, expect = genuine]
-nitpick [card = 5, expect = genuine]
 nitpick [card = 10, expect = genuine]
 oops
 
 lemma "f (a, b) = x"
-nitpick [card = 3, expect = genuine]
-nitpick [card = 10, expect = genuine]
-nitpick [card = 16, expect = genuine]
-nitpick [card = 30, expect = genuine]
+nitpick [card = 12, expect = genuine]
 oops
 
 lemma "f (a, a) = f (c, d)"
-nitpick [card = 4, expect = genuine]
-nitpick [card = 20, expect = genuine]
+nitpick [card = 12, expect = genuine]
 oops
 
 lemma "(x\<Colon>'a) = (\<lambda>a. \<lambda>b. \<lambda>c. if c then a else b) x x True"
-nitpick [card = 2, expect = none]
+nitpick [card = 1\<midarrow>12, expect = none]
 by auto
 
 lemma "\<exists>F. F a b = G a b"
-nitpick [card = 3, expect = none]
+nitpick [card = 2, expect = none]
 by auto
 
 lemma "f = split"
@@ -216,12 +192,10 @@
 
 lemma "(A, B) \<in> R \<or> (\<exists>C. (A, C) \<in> R \<and> (C, B) \<in> R) \<Longrightarrow> 
        A = B \<or> (A, B) \<in> R \<or> (\<exists>C. (A, C) \<in> R \<and> (C, B) \<in> R)"
-nitpick [card = 1\<midarrow>50, expect = none]
+nitpick [card = 1\<midarrow>25, expect = none]
 by auto
 
 lemma "f = (\<lambda>x\<Colon>'a\<times>'b. x)"
-nitpick [card = 3, expect = genuine]
-nitpick [card = 4, expect = genuine]
 nitpick [card = 8, expect = genuine]
 oops
 
@@ -230,30 +204,26 @@
 lemma "x = y"
 nitpick [card 'a = 1, expect = none]
 nitpick [card 'a = 2, expect = genuine]
-nitpick [card 'a = 100, expect = genuine]
-nitpick [card 'a = 1000, expect = genuine]
+nitpick [card 'a = 200, expect = genuine]
 oops
 
 lemma "\<forall>x. x = y"
 nitpick [card 'a = 1, expect = none]
 nitpick [card 'a = 2, expect = genuine]
-nitpick [card 'a = 100, expect = genuine]
-nitpick [card 'a = 1000, expect = genuine]
+nitpick [card 'a = 200, expect = genuine]
 oops
 
 lemma "\<forall>x\<Colon>'a \<Rightarrow> bool. x = y"
 nitpick [card 'a = 1, expect = genuine]
-nitpick [card 'a = 2, expect = genuine]
-nitpick [card 'a = 100, expect = genuine]
-nitpick [card 'a = 1000, expect = genuine]
+nitpick [card 'a = 200, expect = genuine]
 oops
 
 lemma "\<exists>x\<Colon>'a \<Rightarrow> bool. x = y"
-nitpick [card 'a = 1\<midarrow>10, expect = none]
+nitpick [card 'a = 1\<midarrow>20, expect = none]
 by auto
 
 lemma "\<exists>x y\<Colon>'a \<Rightarrow> bool. x = y"
-nitpick [card = 1\<midarrow>40, expect = none]
+nitpick [card = 1\<midarrow>20, expect = none]
 by auto
 
 lemma "\<forall>x. \<exists>y. f x y = f x (g x)"
@@ -261,11 +231,10 @@
 by auto
 
 lemma "\<forall>u. \<exists>v. \<forall>w. \<exists>x. f u v w x = f u (g u) w (h u w)"
-nitpick [card = 1\<midarrow>5, expect = none]
+nitpick [card = 1\<midarrow>4, expect = none]
 by auto
 
 lemma "\<forall>u. \<exists>v. \<forall>w. \<exists>x. f u v w x = f u (g u w) w (h u)"
-nitpick [card = 1\<midarrow>2, expect = genuine]
 nitpick [card = 3, expect = genuine]
 oops
 
@@ -273,7 +242,6 @@
        f u v w x y z = f u (g u) w (h u w) y (k u w y)"
 nitpick [card = 1\<midarrow>2, expect = none]
 nitpick [card = 3, expect = none]
-nitpick [card = 4, expect = none]
 sorry
 
 lemma "\<forall>u. \<exists>v. \<forall>w. \<exists>x. \<forall>y. \<exists>z.
@@ -334,9 +302,6 @@
 
 lemma "(\<exists>x\<Colon>'a. \<forall>y. P x y) \<or> (\<exists>x\<Colon>'a \<times> 'a. \<forall>y. P y x)"
 nitpick [card 'a = 1, expect = genuine]
-nitpick [card 'a = 2, expect = genuine]
-nitpick [card 'a = 3, expect = genuine]
-nitpick [card 'a = 4, expect = genuine]
 nitpick [card 'a = 5, expect = genuine]
 oops
 
@@ -390,8 +355,7 @@
 nitpick [card = 1, expect = genuine]
 nitpick [card = 1, box "('a \<Rightarrow> prop) \<Rightarrow> prop", expect = genuine]
 nitpick [card = 2, expect = genuine]
-nitpick [card = 8, expect = genuine]
-nitpick [card = 10, expect = unknown]
+nitpick [card = 6, expect = genuine]
 oops
 
 lemma "\<And>x. f x y = f x y"
@@ -416,11 +380,7 @@
 
 lemma "x \<equiv> (op \<equiv>) \<Longrightarrow> False"
 nitpick [card = 1, expect = genuine]
-nitpick [card = 2, expect = genuine]
-nitpick [card = 3, expect = genuine]
-nitpick [card = 4, expect = genuine]
-nitpick [card = 5, expect = genuine]
-nitpick [card = 100, expect = genuine]
+nitpick [card = 20, expect = genuine]
 oops
 
 lemma "I = (\<lambda>x. x) \<Longrightarrow> (op \<equiv> x) \<equiv> (\<lambda>y. (x \<equiv> I y))"
@@ -529,7 +489,7 @@
 lemma "x = Ex \<Longrightarrow> False"
 nitpick [card = 1, dont_box, expect = genuine]
 nitpick [card = 2, dont_box, expect = genuine]
-nitpick [card = 8, dont_box, expect = genuine]
+nitpick [card = 6, dont_box, expect = genuine]
 nitpick [card = 10, dont_box, expect = unknown]
 oops
 
@@ -582,8 +542,8 @@
 nitpick [expect = none]
 by auto
 
-lemma "I = (\<lambda>x. x) \<Longrightarrow> (op &) = (\<lambda>x. op & (I x))"
-      "I = (\<lambda>x. x) \<Longrightarrow> (op &) = (\<lambda>x y. x & (I y))"
+lemma "I = (\<lambda>x. x) \<Longrightarrow> (op \<and>) = (\<lambda>x. op \<and> (I x))"
+      "I = (\<lambda>x. x) \<Longrightarrow> (op \<and>) = (\<lambda>x y. x \<and> (I y))"
 nitpick [expect = none]
 by auto
 
@@ -796,7 +756,7 @@
 by auto
 
 lemma "((x, x), (x, x)) \<in> rtrancl {}"
-nitpick [expect = none]
+nitpick [card = 1\<midarrow>5, expect = none]
 by auto
 
 lemma "I = (\<lambda>x. x) \<Longrightarrow> op \<union> = (\<lambda>x. op \<union> (I x))"
@@ -931,9 +891,8 @@
 oops
 
 lemma "P x \<Longrightarrow> P (The P)"
-nitpick [card = 1, expect = none]
 nitpick [card = 1\<midarrow>2, expect = none]
-nitpick [card = 3\<midarrow>5, expect = genuine]
+nitpick [card = 3, expect = genuine]
 nitpick [card = 8, expect = genuine]
 oops
 
--- a/src/HOL/Nitpick_Examples/Datatype_Nits.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Datatype_Nits.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -11,7 +11,8 @@
 imports Main
 begin
 
-nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60 s]
+nitpick_params [max_potential = 0, sat_solver = MiniSat_JNI, max_threads = 1,
+                timeout = 60 s]
 
 primrec rot where
 "rot Nibble0 = Nibble1" | "rot Nibble1 = Nibble2" | "rot Nibble2 = Nibble3" |
@@ -27,9 +28,8 @@
 
 lemma "rot Nibble2 \<noteq> Nibble3"
 nitpick [card = 1, expect = none]
-nitpick [card = 2, expect = genuine]
+nitpick [card = 2, max Nibble4 = 0, expect = genuine]
 nitpick [card = 2, max Nibble2 = 0, expect = none]
-nitpick [card = 2, max Nibble3 = 0, expect = none]
 oops
 
 lemma "(rot ^^ 15) n \<noteq> n"
@@ -53,7 +53,7 @@
 "sn (Pd (_, b)) = b"
 
 lemma "fs (Pd p) = fst p"
-nitpick [card = 20, expect = none]
+nitpick [card = 12, expect = none]
 sorry
 
 lemma "fs (Pd p) = snd p"
@@ -61,7 +61,7 @@
 oops
 
 lemma "sn (Pd p) = snd p"
-nitpick [card = 20, expect = none]
+nitpick [card = 12, expect = none]
 sorry
 
 lemma "sn (Pd p) = fst p"
@@ -69,7 +69,7 @@
 oops
 
 lemma "fs (Pd ((a, b), (c, d))) = (a, b)"
-nitpick [card = 1\<midarrow>12, expect = none]
+nitpick [card = 1\<midarrow>10, expect = none]
 sorry
 
 lemma "fs (Pd ((a, b), (c, d))) = (c, d)"
@@ -82,7 +82,7 @@
 "app (Fn f) x = f x"
 
 lemma "app (Fn g) y = g y"
-nitpick [card = 1\<midarrow>12, expect = none]
+nitpick [card = 1\<midarrow>10, expect = none]
 sorry
 
 lemma "app (Fn g) y = g' y"
--- a/src/HOL/Nitpick_Examples/Hotel_Nits.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Hotel_Nits.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -12,7 +12,8 @@
 imports Main
 begin
 
-nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 120 s]
+nitpick_params [max_potential = 0, sat_solver = MiniSat_JNI, max_threads = 1,
+                timeout = 120 s]
 
 typedecl guest
 typedecl key
--- a/src/HOL/Nitpick_Examples/Manual_Nits.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -8,7 +8,7 @@
 header {* Examples from the Nitpick Manual *}
 
 theory Manual_Nits
-imports Main Coinductive_List RealDef
+imports Main Coinductive_List Quotient_Product RealDef
 begin
 
 chapter {* 3. First Steps *}
@@ -102,6 +102,21 @@
 nitpick [show_datatypes]
 oops
 
+fun my_int_rel where
+"my_int_rel (x, y) (u, v) = (x + v = u + y)"
+
+quotient_type my_int = "nat \<times> nat" / my_int_rel
+by (auto simp add: equivp_def expand_fun_eq)
+
+definition add_raw where
+"add_raw \<equiv> \<lambda>(x, y) (u, v). (x + (u\<Colon>nat), y + (v\<Colon>nat))"
+
+quotient_definition "add\<Colon>my_int \<Rightarrow> my_int \<Rightarrow> my_int" is add_raw
+
+lemma "add x y = add x x"
+nitpick [show_datatypes]
+oops
+
 record point =
   Xcoord :: int
   Ycoord :: int
--- a/src/HOL/Nitpick_Examples/Mini_Nits.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Mini_Nits.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -15,11 +15,9 @@
 exception FAIL
 
 (* int -> term -> string *)
-fun minipick 0 _ = "none"
-  | minipick n t =
-    case minipick (n - 1) t of
-      "none" => Minipick.pick_nits_in_term @{context} (K n) t
-    | outcome_code => outcome_code
+fun minipick n t =
+  map (fn k => Minipick.kodkod_problem_from_term @{context} (K k) t) (1 upto n)
+  |> Minipick.solve_any_kodkod_problem @{theory}
 (* int -> term -> bool *)
 fun none n t = (minipick n t = "none" orelse raise FAIL)
 fun genuine n t = (minipick n t = "genuine" orelse raise FAIL)
@@ -87,11 +85,11 @@
 ML {* genuine 2 @{prop "(a\<Colon>'a \<times> 'a \<times> 'a, b) \<noteq> (b, a)"} *}
 ML {* none 5 @{prop "\<exists>a b\<Colon>'a \<Rightarrow> 'a. (a, b) = (b, a)"} *}
 ML {* genuine 1 @{prop "(a\<Colon>'a \<Rightarrow> 'a, b) \<noteq> (b, a)"} *}
-ML {* none 8 @{prop "fst (a, b) = a"} *}
+ML {* none 5 @{prop "fst (a, b) = a"} *}
 ML {* none 1 @{prop "fst (a, b) = b"} *}
 ML {* genuine 2 @{prop "fst (a, b) = b"} *}
 ML {* genuine 2 @{prop "fst (a, b) \<noteq> b"} *}
-ML {* none 8 @{prop "snd (a, b) = b"} *}
+ML {* none 5 @{prop "snd (a, b) = b"} *}
 ML {* none 1 @{prop "snd (a, b) = a"} *}
 ML {* genuine 2 @{prop "snd (a, b) = a"} *}
 ML {* genuine 2 @{prop "snd (a, b) \<noteq> a"} *}
--- a/src/HOL/Nitpick_Examples/Pattern_Nits.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Pattern_Nits.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -11,8 +11,8 @@
 imports Main
 begin
 
-nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60 s,
-                card = 14]
+nitpick_params [card = 10, max_potential = 0, sat_solver = MiniSat_JNI,
+                max_threads = 1, timeout = 60 s]
 
 lemma "x = (case u of () \<Rightarrow> y)"
 nitpick [expect = genuine]
@@ -132,7 +132,7 @@
 nitpick [expect = genuine]
 oops
 
-lemma "\<exists>y a b zs. x = (y # Some (a, b) # zs)"
+lemma "\<exists>y a b zs. x = y # Some (a, b) # zs"
 nitpick [expect = genuine]
 oops
 
--- a/src/HOL/Nitpick_Examples/Record_Nits.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Record_Nits.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -11,7 +11,8 @@
 imports Main
 begin
 
-nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60 s]
+nitpick_params [max_potential = 0, sat_solver = MiniSat_JNI, max_threads = 1,
+                timeout = 60 s]
 
 record point2d =
   xc :: int
--- a/src/HOL/Nitpick_Examples/Refute_Nits.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Refute_Nits.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -11,7 +11,8 @@
 imports Main
 begin
 
-nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60 s]
+nitpick_params [max_potential = 0, sat_solver = MiniSat_JNI, max_threads = 1,
+                timeout = 60 s]
 
 lemma "P \<and> Q"
 apply (rule conjI)
@@ -174,14 +175,14 @@
 nitpick [expect = genuine]
 oops
 
-text {* The "Drinker's theorem" ... *}
+text {* The ``Drinker's theorem'' *}
 
 lemma "\<exists>x. f x = g x \<longrightarrow> f = g"
 nitpick [expect = none]
 apply (auto simp add: ext)
 done
 
-text {* ... and an incorrect version of it *}
+text {* And an incorrect version of it *}
 
 lemma "(\<exists>x. f x = g x) \<longrightarrow> f = g"
 nitpick [expect = genuine]
@@ -241,7 +242,7 @@
 nitpick [expect = genuine]
 oops
 
-text {* "The transitive closure 'T' of an arbitrary relation 'P' is non-empty." *}
+text {* ``The transitive closure of an arbitrary relation is non-empty.'' *}
 
 constdefs
 "trans" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
@@ -255,7 +256,7 @@
 nitpick [expect = genuine]
 oops
 
-text {* "The union of transitive closures is equal to the transitive closure of unions." *}
+text {* ``The union of transitive closures is equal to the transitive closure of unions.'' *}
 
 lemma "(\<forall>x y. (P x y \<or> R x y) \<longrightarrow> T x y) \<longrightarrow> trans T \<longrightarrow> (\<forall>Q. (\<forall>x y. (P x y \<or> R x y) \<longrightarrow> Q x y) \<longrightarrow> trans Q \<longrightarrow> subset T Q)
  \<longrightarrow> trans_closure TP P
@@ -264,19 +265,19 @@
 nitpick [expect = genuine]
 oops
 
-text {* "Every surjective function is invertible." *}
+text {* ``Every surjective function is invertible.'' *}
 
 lemma "(\<forall>y. \<exists>x. y = f x) \<longrightarrow> (\<exists>g. \<forall>x. g (f x) = x)"
 nitpick [expect = genuine]
 oops
 
-text {* "Every invertible function is surjective." *}
+text {* ``Every invertible function is surjective.'' *}
 
 lemma "(\<exists>g. \<forall>x. g (f x) = x) \<longrightarrow> (\<forall>y. \<exists>x. y = f x)"
 nitpick [expect = genuine]
 oops
 
-text {* Every point is a fixed point of some function. *}
+text {* ``Every point is a fixed point of some function.'' *}
 
 lemma "\<exists>f. f x = x"
 nitpick [card = 1\<midarrow>7, expect = none]
@@ -284,21 +285,21 @@
 apply simp
 done
 
-text {* Axiom of Choice: first an incorrect version ... *}
+text {* Axiom of Choice: first an incorrect version *}
 
 lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>!f. \<forall>x. P x (f x))"
 nitpick [expect = genuine]
 oops
 
-text {* ... and now two correct ones *}
+text {* And now two correct ones *}
 
 lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>f. \<forall>x. P x (f x))"
-nitpick [card = 1-5, expect = none]
+nitpick [card = 1-4, expect = none]
 apply (simp add: choice)
 done
 
 lemma "(\<forall>x. \<exists>!y. P x y) \<longrightarrow> (\<exists>!f. \<forall>x. P x (f x))"
-nitpick [card = 1-4, expect = none]
+nitpick [card = 1-3, expect = none]
 apply auto
  apply (simp add: ex1_implies_ex choice)
 apply (fast intro: ext)
@@ -807,12 +808,12 @@
 oops
 
 lemma "list_rec nil cons [] = nil"
-nitpick [expect = none]
+nitpick [card = 1\<midarrow>5, expect = none]
 apply simp
 done
 
 lemma "list_rec nil cons (x#xs) = cons x xs (list_rec nil cons xs)"
-nitpick [expect = none]
+nitpick [card = 1\<midarrow>5, expect = none]
 apply simp
 done
 
@@ -923,12 +924,12 @@
 oops
 
 lemma "aexp_bexp_rec_1 number ite equal (Number x) = number x"
-nitpick [card = 1\<midarrow>4, expect = none]
+nitpick [card = 1\<midarrow>3, expect = none]
 apply simp
 done
 
 lemma "aexp_bexp_rec_1 number ite equal (ITE x y z) = ite x y z (aexp_bexp_rec_2 number ite equal x) (aexp_bexp_rec_1 number ite equal y) (aexp_bexp_rec_1 number ite equal z)"
-nitpick [card = 1\<midarrow>4, expect = none]
+nitpick [card = 1\<midarrow>3, expect = none]
 apply simp
 done
 
@@ -941,7 +942,7 @@
 oops
 
 lemma "aexp_bexp_rec_2 number ite equal (Equal x y) = equal x y (aexp_bexp_rec_1 number ite equal x) (aexp_bexp_rec_1 number ite equal y)"
-nitpick [card = 1\<midarrow>4, expect = none]
+nitpick [card = 1\<midarrow>3, expect = none]
 apply simp
 done
 
@@ -1001,32 +1002,32 @@
 oops
 
 lemma "X_Y_rec_1 a b c d e f A = a"
-nitpick [expect = none]
+nitpick [card = 1\<midarrow>5, expect = none]
 apply simp
 done
 
 lemma "X_Y_rec_1 a b c d e f (B x) = b x (X_Y_rec_1 a b c d e f x)"
-nitpick [expect = none]
+nitpick [card = 1\<midarrow>5, expect = none]
 apply simp
 done
 
 lemma "X_Y_rec_1 a b c d e f (C y) = c y (X_Y_rec_2 a b c d e f y)"
-nitpick [expect = none]
+nitpick [card = 1\<midarrow>5, expect = none]
 apply simp
 done
 
 lemma "X_Y_rec_2 a b c d e f (D x) = d x (X_Y_rec_1 a b c d e f x)"
-nitpick [expect = none]
+nitpick [card = 1\<midarrow>5, expect = none]
 apply simp
 done
 
 lemma "X_Y_rec_2 a b c d e f (E y) = e y (X_Y_rec_2 a b c d e f y)"
-nitpick [expect = none]
+nitpick [card = 1\<midarrow>5, expect = none]
 apply simp
 done
 
 lemma "X_Y_rec_2 a b c d e f F = f"
-nitpick [expect = none]
+nitpick [card = 1\<midarrow>5, expect = none]
 apply simp
 done
 
@@ -1057,12 +1058,12 @@
 oops
 
 lemma "XOpt_rec_1 cx dx n1 s1 n2 s2 (CX x) = cx x (XOpt_rec_2 cx dx n1 s1 n2 s2 x)"
-nitpick [card = 1\<midarrow>6, expect = none]
+nitpick [card = 1\<midarrow>5, expect = none]
 apply simp
 done
 
 lemma "XOpt_rec_1 cx dx n1 s1 n2 s2 (DX x) = dx x (\<lambda>b. XOpt_rec_3 cx dx n1 s1 n2 s2 (x b))"
-nitpick [card = 1\<midarrow>4, expect = none]
+nitpick [card = 1\<midarrow>3, expect = none]
 apply simp
 done
 
@@ -1150,17 +1151,17 @@
 oops
 
 lemma "Trie_rec_1 tr nil cons (TR x) = tr x (Trie_rec_2 tr nil cons x)"
-nitpick [card = 1\<midarrow>6, expect = none]
+nitpick [card = 1\<midarrow>4, expect = none]
 apply simp
 done
 
 lemma "Trie_rec_2 tr nil cons [] = nil"
-nitpick [card = 1\<midarrow>6, expect = none]
+nitpick [card = 1\<midarrow>4, expect = none]
 apply simp
 done
 
 lemma "Trie_rec_2 tr nil cons (x#xs) = cons x xs (Trie_rec_1 tr nil cons x) (Trie_rec_2 tr nil cons xs)"
-nitpick [card = 1\<midarrow>6, expect = none]
+nitpick [card = 1\<midarrow>4, expect = none]
 apply simp
 done
 
@@ -1365,15 +1366,15 @@
 oops
 
 lemma "f (lfp f) = lfp f"
-nitpick [expect = genuine]
+nitpick [card = 2, expect = genuine]
 oops
 
 lemma "f (gfp f) = gfp f"
-nitpick [expect = genuine]
+nitpick [card = 2, expect = genuine]
 oops
 
 lemma "lfp f = gfp f"
-nitpick [expect = genuine]
+nitpick [card = 2, expect = genuine]
 oops
 
 subsubsection {* Axiomatic Type Classes and Overloading *}
--- a/src/HOL/Nitpick_Examples/Special_Nits.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Special_Nits.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -11,8 +11,8 @@
 imports Main
 begin
 
-nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60 s,
-                card = 4]
+nitpick_params [card = 4, sat_solver = MiniSat_JNI, max_threads = 1,
+                timeout = 60 s]
 
 fun f1 :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" where
 "f1 a b c d e = a + b + c + d + e"
--- a/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -11,8 +11,8 @@
 imports Main Rational
 begin
 
-nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60 s,
-                card = 1\<midarrow>4]
+nitpick_params [card = 1\<midarrow>4, sat_solver = MiniSat_JNI, max_threads = 1,
+                timeout = 60 s]
 
 typedef three = "{0\<Colon>nat, 1, 2}"
 by blast
@@ -138,7 +138,8 @@
 by (rule Inl_Rep_not_Inr_Rep)
 
 lemma "Abs_Sum (Rep_Sum a) = a"
-nitpick [card = 1\<midarrow>2, timeout = 60 s, expect = none]
+nitpick [card = 1, expect = none]
+nitpick [card = 2, expect = none]
 by (rule Rep_Sum_inverse)
 
 lemma "0::nat \<equiv> Abs_Nat Zero_Rep"
--- a/src/HOL/Prolog/Func.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Prolog/Func.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -10,31 +10,28 @@
 
 typedecl tm
 
-consts
-  abs     :: "(tm => tm) => tm"
-  app     :: "tm => tm => tm"
+axiomatization
+  abs     :: "(tm \<Rightarrow> tm) \<Rightarrow> tm" and
+  app     :: "tm \<Rightarrow> tm \<Rightarrow> tm" and
 
-  cond    :: "tm => tm => tm => tm"
-  "fix"   :: "(tm => tm) => tm"
-
-  true    :: tm
-  false   :: tm
-  "and"   :: "tm => tm => tm"       (infixr "and" 999)
-  eq      :: "tm => tm => tm"       (infixr "eq" 999)
+  cond    :: "tm \<Rightarrow> tm \<Rightarrow> tm \<Rightarrow> tm" and
+  "fix"   :: "(tm \<Rightarrow> tm) \<Rightarrow> tm" and
 
-  Z       :: tm                     ("Z")
-  S       :: "tm => tm"
-(*
-        "++", "--",
-        "**"    :: tm => tm => tm       (infixr 999)
-*)
-        eval    :: "[tm, tm] => bool"
+  true    :: tm and
+  false   :: tm and
+  "and"   :: "tm \<Rightarrow> tm \<Rightarrow> tm"       (infixr "and" 999) and
+  eq      :: "tm \<Rightarrow> tm \<Rightarrow> tm"       (infixr "eq" 999) and
+
+  Z       :: tm                     ("Z") and
+  S       :: "tm \<Rightarrow> tm" and
 
-instance tm :: plus ..
-instance tm :: minus ..
-instance tm :: times ..
+  plus    :: "tm \<Rightarrow> tm \<Rightarrow> tm"       (infixl "+" 65) and
+  minus   :: "tm \<Rightarrow> tm \<Rightarrow> tm"       (infixl "-" 65) and
+  times   :: "tm \<Rightarrow> tm \<Rightarrow> tm"       (infixl "*" 70) and
 
-axioms   eval: "
+  eval    :: "tm \<Rightarrow> tm \<Rightarrow> bool" where
+
+eval: "
 
 eval (abs RR) (abs RR)..
 eval (app F X) V :- eval F (abs R) & eval X U & eval (R U) V..
@@ -59,7 +56,6 @@
 eval ( Z    * M) Z..
 eval ((S N) * M) K :- eval (N * M) L & eval (L + M) K"
 
-
 lemmas prog_Func = eval
 
 lemma "eval ((S (S Z)) + (S Z)) ?X"
--- a/src/HOL/Random.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Random.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -168,6 +168,7 @@
 hide (open) type seed
 hide (open) const inc_shift minus_shift log "next" split_seed
   iterate range select pick select_weight
+hide (open) fact range_def
 
 no_notation fcomp (infixl "o>" 60)
 no_notation scomp (infixl "o\<rightarrow>" 60)
--- a/src/HOL/Tools/Datatype/datatype_case.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Datatype/datatype_case.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -310,6 +310,8 @@
 fun case_tr err tab_of ctxt [t, u] =
     let
       val thy = ProofContext.theory_of ctxt;
+      val intern_const_syntax = Consts.intern_syntax (Sign.consts_of thy);
+
       (* replace occurrences of dummy_pattern by distinct variables *)
       (* internalize constant names                                 *)
       fun prep_pat ((c as Const (@{syntax_const "_constrain"}, _)) $ t $ tT) used =
@@ -319,9 +321,7 @@
             let val x = Name.variant used "x"
             in (Free (x, T), x :: used) end
         | prep_pat (Const (s, T)) used =
-            (case try (unprefix Syntax.constN) s of
-               SOME c => (Const (c, T), used)
-             | NONE => (Const (Sign.intern_const thy s, T), used))
+            (Const (intern_const_syntax s, T), used)
         | prep_pat (v as Free (s, T)) used =
             let val s' = Sign.intern_const thy s
             in
@@ -422,8 +422,7 @@
   | _ => NONE;
 
 val dest_case = gen_dest_case (try (dest_Const #> fst)) fastype_of;
-val dest_case' = gen_dest_case
-  (try (dest_Const #> fst #> unprefix Syntax.constN)) (K dummyT);
+val dest_case' = gen_dest_case (try (dest_Const #> fst #> Syntax.unmark_const)) (K dummyT);
 
 
 (* destruct nested patterns *)
@@ -455,14 +454,13 @@
 fun case_tr' tab_of cname ctxt ts =
   let
     val thy = ProofContext.theory_of ctxt;
-    val consts = ProofContext.consts_of ctxt;
     fun mk_clause (pat, rhs) =
       let val xs = Term.add_frees pat []
       in
         Syntax.const @{syntax_const "_case1"} $
           map_aterms
             (fn Free p => Syntax.mark_boundT p
-              | Const (s, _) => Const (Consts.extern_early consts s, dummyT)
+              | Const (s, _) => Syntax.const (Syntax.mark_const s)
               | t => t) pat $
           map_aterms
             (fn x as Free (s, T) =>
--- a/src/HOL/Tools/Datatype/datatype_data.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Datatype/datatype_data.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -223,7 +223,7 @@
 fun add_case_tr' case_names thy =
   Sign.add_advanced_trfuns ([], [],
     map (fn case_name =>
-      let val case_name' = Sign.const_syntax_name thy case_name
+      let val case_name' = Syntax.mark_const case_name
       in (case_name', Datatype_Case.case_tr' info_of_case case_name')
       end) case_names, []) thy;
 
--- a/src/HOL/Tools/Function/size.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Function/size.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -25,7 +25,7 @@
 
 val lookup_size = SizeData.get #> Symtab.lookup;
 
-fun plus (t1, t2) = Const (@{const_name Algebras.plus},
+fun plus (t1, t2) = Const (@{const_name Groups.plus},
   HOLogic.natT --> HOLogic.natT --> HOLogic.natT) $ t1 $ t2;
 
 fun size_of_type f g h (T as Type (s, Ts)) =
--- a/src/HOL/Tools/Nitpick/kodkod.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Nitpick/kodkod.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -465,7 +465,7 @@
   | arity_of_rel_expr (Intersect (r1, _)) = arity_of_rel_expr r1
   | arity_of_rel_expr (Product (r1, r2)) = sum_arities_of_rel_exprs r1 r2
   | arity_of_rel_expr (IfNo (r1, _)) = arity_of_rel_expr r1
-  | arity_of_rel_expr (Project (r, is)) = length is
+  | arity_of_rel_expr (Project (_, is)) = length is
   | arity_of_rel_expr (Join (r1, r2)) = sum_arities_of_rel_exprs r1 r2 - 2
   | arity_of_rel_expr (Closure _) = 2
   | arity_of_rel_expr (ReflexiveClosure _) = 2
@@ -590,8 +590,8 @@
         (case tuple_set of
            TupleUnion (ts1, ts2) => sub ts1 prec ^ " + " ^ sub ts2 (prec + 1)
          | TupleDifference (ts1, ts2) =>
-           sub ts1 prec ^ " - " ^ sub ts1 (prec + 1)
-         | TupleIntersect (ts1, ts2) => sub ts1 prec ^ " & " ^ sub ts1 prec
+           sub ts1 prec ^ " - " ^ sub ts2 (prec + 1)
+         | TupleIntersect (ts1, ts2) => sub ts1 prec ^ " & " ^ sub ts2 prec
          | TupleProduct (ts1, ts2) => sub ts1 prec ^ "->" ^ sub ts2 prec
          | TupleProject (ts, c) => sub ts prec ^ "[" ^ string_of_int c ^ "]"
          | TupleSet ts => "{" ^ commas (map string_for_tuple ts) ^ "}"
--- a/src/HOL/Tools/Nitpick/minipick.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Nitpick/minipick.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -19,7 +19,9 @@
   val true_atom : Kodkod.rel_expr
   val formula_from_atom : Kodkod.rel_expr -> Kodkod.formula
   val atom_from_formula : Kodkod.formula -> Kodkod.rel_expr
-  val pick_nits_in_term : Proof.context -> (typ -> int) -> term -> string
+  val kodkod_problem_from_term :
+    Proof.context -> (typ -> int) -> term -> Kodkod.problem
+  val solve_any_kodkod_problem : theory -> Kodkod.problem list -> string
 end;
 
 structure Minipick : MINIPICK =
@@ -128,7 +130,7 @@
          All (decls_for SRep card Ts T, to_F (T :: Ts) t')
        | (t0 as Const (@{const_name All}, _)) $ t1 =>
          to_F Ts (t0 $ eta_expand Ts t1 1)
-       | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
+       | Const (@{const_name Ex}, _) $ Abs (_, T, t') =>
          Exist (decls_for SRep card Ts T, to_F (T :: Ts) t')
        | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
          to_F Ts (t0 $ eta_expand Ts t1 1)
@@ -234,7 +236,7 @@
        | Free (x as (_, T)) =>
          Rel (arity_of RRep card T, find_index (curry (op =) x) frees)
        | Term.Var _ => raise NOT_SUPPORTED "schematic variables"
-       | Bound j => raise SAME ()
+       | Bound _ => raise SAME ()
        | Abs (_, T, t') =>
          (case fastype_of1 (T :: Ts, t') of
             @{typ bool} => Comprehension (decls_for SRep card Ts T,
@@ -251,11 +253,8 @@
             let val T2 = fastype_of1 (Ts, t2) in
               case arity_of SRep card T2 of
                 1 => Join (to_S_rep Ts t2, to_R_rep Ts t1)
-              | n =>
-                let
-                  val arity2 = arity_of SRep card T2
-                  val res_arity = arity_of RRep card T
-                in
+              | arity2 =>
+                let val res_arity = arity_of RRep card T in
                   Project (Intersect
                       (Product (to_S_rep Ts t2,
                                 atom_schema_of RRep card T
@@ -290,11 +289,10 @@
 fun declarative_axiom_for_free card i (_, T) =
   declarative_axiom_for_rel_expr card [] T (Rel (arity_of RRep card T, i))
 
-(* Proof.context -> (typ -> int) -> term -> string *)
-fun pick_nits_in_term ctxt raw_card t =
+(* Proof.context -> (typ -> int) -> term -> problem *)
+fun kodkod_problem_from_term ctxt raw_card t =
   let
     val thy = ProofContext.theory_of ctxt
-    val {overlord, ...} = Nitpick_Isar.default_params thy []
     (* typ -> int *)
     fun card (Type ("fun", [T1, T2])) = reasonable_power (card T2) (card T1)
       | card (Type ("*", [T1, T2])) = card T1 * card T2
@@ -309,11 +307,19 @@
     val formula = kodkod_formula_from_term ctxt card frees neg_t
                   |> fold_rev (curry And) declarative_axioms
     val univ_card = univ_card 0 0 0 bounds formula
-    val problem =
-      {comment = "", settings = [], univ_card = univ_card, tuple_assigns = [],
-       bounds = bounds, int_bounds = [], expr_assigns = [], formula = formula}
   in
-    case solve_any_problem overlord NONE 0 1 [problem] of
+    {comment = "", settings = [], univ_card = univ_card, tuple_assigns = [],
+     bounds = bounds, int_bounds = [], expr_assigns = [], formula = formula}
+  end
+
+(* theory -> problem list -> string *)
+fun solve_any_kodkod_problem thy problems =
+  let
+    val {overlord, ...} = Nitpick_Isar.default_params thy []
+    val max_threads = 1
+    val max_solutions = 1
+  in
+    case solve_any_problem overlord NONE max_threads max_solutions problems of
       NotInstalled => "unknown"
     | Normal ([], _) => "none"
     | Normal _ => "genuine"
@@ -321,7 +327,5 @@
     | Interrupted _ => "unknown"
     | Error (s, _) => error ("Kodkod error: " ^ s)
   end
-  handle NOT_SUPPORTED details =>
-         (warning ("Unsupported case: " ^ details ^ "."); "unknown")
 
 end;
--- a/src/HOL/Tools/Nitpick/nitpick.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -201,13 +201,13 @@
             error "You must import the theory \"Nitpick\" to use Nitpick"
 *)
     val {cards_assigns, maxes_assigns, iters_assigns, bitss, bisim_depths,
-         boxes, monos, stds, wfs, sat_solver, blocking, falsify, debug, verbose,
-         overlord, user_axioms, assms, merge_type_vars, binary_ints,
-         destroy_constrs, specialize, skolemize, star_linear_preds, uncurry,
-         fast_descrs, peephole_optim, tac_timeout, sym_break, sharing_depth,
-         flatten_props, max_threads, show_skolems, show_datatypes, show_consts,
-         evals, formats, max_potential, max_genuine, check_potential,
-         check_genuine, batch_size, ...} =
+         boxes, monos, stds, wfs, sat_solver, falsify, debug, verbose, overlord,
+         user_axioms, assms, merge_type_vars, binary_ints, destroy_constrs,
+         specialize, skolemize, star_linear_preds, uncurry, fast_descrs,
+         peephole_optim, tac_timeout, sym_break, sharing_depth, flatten_props,
+         max_threads, show_skolems, show_datatypes, show_consts, evals, formats,
+         max_potential, max_genuine, check_potential, check_genuine, batch_size,
+         ...} =
       params
     val state_ref = Unsynchronized.ref state
     (* Pretty.T -> unit *)
@@ -227,7 +227,6 @@
     (* (unit -> string) -> unit *)
     val print_m = pprint_m o K o plazy
     val print_v = pprint_v o K o plazy
-    val print_d = pprint_d o K o plazy
 
     (* unit -> unit *)
     fun check_deadline () =
@@ -489,9 +488,9 @@
         val core_u = rename_vars_in_nut pool rel_table core_u
         val def_us = map (rename_vars_in_nut pool rel_table) def_us
         val nondef_us = map (rename_vars_in_nut pool rel_table) nondef_us
-        val core_f = kodkod_formula_from_nut bits ofs kk core_u
-        val def_fs = map (kodkod_formula_from_nut bits ofs kk) def_us
-        val nondef_fs = map (kodkod_formula_from_nut bits ofs kk) nondef_us
+        val core_f = kodkod_formula_from_nut ofs kk core_u
+        val def_fs = map (kodkod_formula_from_nut ofs kk) def_us
+        val nondef_fs = map (kodkod_formula_from_nut ofs kk) nondef_us
         val formula = fold (fold s_and) [def_fs, nondef_fs] core_f
         val comment = (if unsound then "unsound" else "sound") ^ "\n" ^
                       PrintMode.setmp [] multiline_string_for_scope scope
@@ -535,9 +534,8 @@
       in
         SOME ({comment = comment, settings = settings, univ_card = univ_card,
                tuple_assigns = [], bounds = bounds,
-               int_bounds =
-                   if bits = 0 then sequential_int_bounds univ_card
-                   else pow_of_two_int_bounds bits main_j0 univ_card,
+               int_bounds = if bits = 0 then sequential_int_bounds univ_card
+                            else pow_of_two_int_bounds bits main_j0,
                expr_assigns = [], formula = formula},
               {free_names = free_names, sel_names = sel_names,
                nonsel_names = nonsel_names, rel_table = rel_table,
@@ -562,17 +560,13 @@
                                                     else "genuine") ^
                                    " component of scope."));
                 NONE)
-           | TOO_SMALL (loc, msg) =>
+           | TOO_SMALL (_, msg) =>
              (print_v (fn () => ("Limit reached: " ^ msg ^
                                  ". Skipping " ^ (if unsound then "potential"
                                                   else "genuine") ^
                                  " component of scope."));
               NONE)
 
-    (* int -> (''a * int list list) list -> ''a -> KK.tuple_set *)
-    fun tuple_set_for_rel univ_card =
-      KK.TupleSet o map (kk_tuple debug univ_card) o the oo AList.lookup (op =)
-
     val das_wort_model =
       (if falsify then "counterexample" else "model")
       |> not standard ? prefix "nonstandard "
@@ -809,8 +803,7 @@
             ()
         (* scope * bool -> rich_problem list * bool
            -> rich_problem list * bool *)
-        fun add_problem_for_scope (scope as {datatypes, ...}, unsound)
-                                  (problems, donno) =
+        fun add_problem_for_scope (scope, unsound) (problems, donno) =
           (check_deadline ();
            case problem_for_scope unsound scope of
              SOME problem =>
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -65,6 +65,7 @@
   val conjuncts_of : term -> term list
   val disjuncts_of : term -> term list
   val unarize_and_unbox_type : typ -> typ
+  val unarize_unbox_and_uniterize_type : typ -> typ
   val string_for_type : Proof.context -> typ -> string
   val prefix_name : string -> string -> string
   val shortest_name : string -> string
@@ -81,6 +82,7 @@
   val is_lfp_iterator_type : typ -> bool
   val is_gfp_iterator_type : typ -> bool
   val is_fp_iterator_type : typ -> bool
+  val is_iterator_type : typ -> bool
   val is_boolean_type : typ -> bool
   val is_integer_type : typ -> bool
   val is_bit_type : typ -> bool
@@ -261,7 +263,6 @@
 val set_prefix = nitpick_prefix ^ "set" ^ name_sep
 val lfp_iterator_prefix = nitpick_prefix ^ "lfpit" ^ name_sep
 val gfp_iterator_prefix = nitpick_prefix ^ "gfpit" ^ name_sep
-val nwf_prefix = nitpick_prefix ^ "nwf" ^ name_sep
 val unrolled_prefix = nitpick_prefix ^ "unroll" ^ name_sep
 val base_prefix = nitpick_prefix ^ "base" ^ name_sep
 val step_prefix = nitpick_prefix ^ "step" ^ name_sep
@@ -306,7 +307,7 @@
     if t0 = conn_t then strip_connective t0 t2 @ strip_connective t0 t1 else [t]
   | strip_connective _ t = [t]
 (* term -> term list * term *)
-fun strip_any_connective (t as (t0 $ t1 $ t2)) =
+fun strip_any_connective (t as (t0 $ _ $ _)) =
     if t0 = @{const "op &"} orelse t0 = @{const "op |"} then
       (strip_connective t0 t, t0)
     else
@@ -389,7 +390,6 @@
 (* typ -> typ *)
 fun unarize_type @{typ "unsigned_bit word"} = nat_T
   | unarize_type @{typ "signed_bit word"} = int_T
-  | unarize_type @{typ bisim_iterator} = nat_T
   | unarize_type (Type (s, Ts as _ :: _)) = Type (s, map unarize_type Ts)
   | unarize_type T = T
 fun unarize_and_unbox_type (Type (@{type_name fun_box}, Ts)) =
@@ -398,10 +398,14 @@
     Type ("*", map unarize_and_unbox_type Ts)
   | unarize_and_unbox_type @{typ "unsigned_bit word"} = nat_T
   | unarize_and_unbox_type @{typ "signed_bit word"} = int_T
-  | unarize_and_unbox_type @{typ bisim_iterator} = nat_T
   | unarize_and_unbox_type (Type (s, Ts as _ :: _)) =
     Type (s, map unarize_and_unbox_type Ts)
   | unarize_and_unbox_type T = T
+fun uniterize_type (Type (s, Ts as _ :: _)) = Type (s, map uniterize_type Ts)
+  | uniterize_type @{typ bisim_iterator} = nat_T
+  | uniterize_type T = T
+val unarize_unbox_and_uniterize_type = uniterize_type o unarize_and_unbox_type
+
 (* Proof.context -> typ -> string *)
 fun string_for_type ctxt = Syntax.string_of_typ ctxt o unarize_and_unbox_type
 
@@ -436,7 +440,7 @@
 fun term_match thy (Const x1, Const x2) = const_match thy (x1, x2)
   | term_match thy (Free (s1, T1), Free (s2, T2)) =
     const_match thy ((shortest_name s1, T1), (shortest_name s2, T2))
-  | term_match thy (t1, t2) = t1 aconv t2
+  | term_match _ (t1, t2) = t1 aconv t2
 
 (* typ -> bool *)
 fun is_TFree (TFree _) = true
@@ -455,14 +459,14 @@
 fun is_gfp_iterator_type (Type (s, _)) = String.isPrefix gfp_iterator_prefix s
   | is_gfp_iterator_type _ = false
 val is_fp_iterator_type = is_lfp_iterator_type orf is_gfp_iterator_type
+fun is_iterator_type T =
+  (T = @{typ bisim_iterator} orelse is_fp_iterator_type T)
 fun is_boolean_type T = (T = prop_T orelse T = bool_T)
 fun is_integer_type T = (T = nat_T orelse T = int_T)
 fun is_bit_type T = (T = @{typ unsigned_bit} orelse T = @{typ signed_bit})
 fun is_word_type (Type (@{type_name word}, _)) = true
   | is_word_type _ = false
-fun is_integer_like_type T =
-  is_fp_iterator_type T orelse is_integer_type T orelse is_word_type T orelse
-  T = @{typ bisim_iterator}
+val is_integer_like_type = is_iterator_type orf is_integer_type orf is_word_type
 val is_record_type = not o null o Record.dest_recTs
 (* theory -> typ -> bool *)
 fun is_frac_type thy (Type (s, [])) =
@@ -593,10 +597,10 @@
 fun unregister_codatatype co_T = register_codatatype co_T "" []
 
 (* theory -> typ -> bool *)
-fun is_quot_type _ (Type ("IntEx.my_int", _)) = true (* FIXME *)
-  | is_quot_type _ (Type ("FSet.fset", _)) = true
+fun is_quot_type thy (Type (s, _)) =
+    is_some (Quotient_Info.quotdata_lookup_raw thy s)
   | is_quot_type _ _ = false
-fun is_codatatype thy (T as Type (s, _)) =
+fun is_codatatype thy (Type (s, _)) =
     not (null (AList.lookup (op =) (#codatatypes (Data.get thy)) s
                |> Option.map snd |> these))
   | is_codatatype _ _ = false
@@ -630,7 +634,7 @@
   end
   handle TYPE _ => []
 (* styp -> bool *)
-fun is_record_constr (x as (s, T)) =
+fun is_record_constr (s, T) =
   String.isSuffix Record.extN s andalso
   let val dataT = body_type T in
     is_record_type dataT andalso
@@ -662,11 +666,13 @@
      | NONE => false)
   | is_rep_fun _ _ = false
 (* Proof.context -> styp -> bool *)
-fun is_quot_abs_fun _ ("IntEx.abs_my_int", _) = true
-  | is_quot_abs_fun _ ("FSet.abs_fset", _) = true
+fun is_quot_abs_fun ctxt (x as (_, Type ("fun", [_, Type (s', _)]))) =
+    (try (Quotient_Term.absrep_const_chk Quotient_Term.AbsF ctxt) s'
+     = SOME (Const x))
   | is_quot_abs_fun _ _ = false
-fun is_quot_rep_fun _ ("IntEx.rep_my_int", _) = true
-  | is_quot_rep_fun _ ("FSet.rep_fset", _) = true
+fun is_quot_rep_fun ctxt (x as (_, Type ("fun", [Type (s', _), _]))) =
+    (try (Quotient_Term.absrep_const_chk Quotient_Term.RepF ctxt) s'
+     = SOME (Const x))
   | is_quot_rep_fun _ _ = false
 
 (* theory -> styp -> styp *)
@@ -676,18 +682,16 @@
      | NONE => raise TERM ("Nitpick_HOL.mate_of_rep_fun", [Const x]))
   | mate_of_rep_fun _ x = raise TERM ("Nitpick_HOL.mate_of_rep_fun", [Const x])
 (* theory -> typ -> typ *)
-fun rep_type_for_quot_type _ (Type ("IntEx.my_int", [])) = @{typ "nat * nat"}
-  | rep_type_for_quot_type _ (Type ("FSet.fset", [T])) =
-    Type (@{type_name list}, [T])
-  | rep_type_for_quot_type _ T =
-    raise TYPE ("Nitpick_HOL.rep_type_for_quot_type", [T], [])
+fun rep_type_for_quot_type thy (T as Type (s, _)) =
+  let val {qtyp, rtyp, ...} = Quotient_Info.quotdata_lookup thy s in
+    instantiate_type thy qtyp T rtyp
+  end
 (* theory -> typ -> term *)
-fun equiv_relation_for_quot_type _ (Type ("IntEx.my_int", [])) =
-    Const ("IntEx.intrel", @{typ "(nat * nat) => (nat * nat) => bool"})
-  | equiv_relation_for_quot_type _ (Type ("FSet.fset", [T])) =
-    Const ("FSet.list_eq",
-           Type (@{type_name list}, [T]) --> Type (@{type_name list}, [T])
-           --> bool_T)
+fun equiv_relation_for_quot_type thy (Type (s, Ts)) =
+    let
+      val {qtyp, equiv_rel, ...} = Quotient_Info.quotdata_lookup thy s
+      val Ts' = qtyp |> dest_Type |> snd
+    in subst_atomic_types (Ts' ~~ Ts) equiv_rel end
   | equiv_relation_for_quot_type _ T =
     raise TYPE ("Nitpick_HOL.equiv_relation_for_quot_type", [T], [])
 
@@ -706,7 +710,7 @@
   member (op =) [@{const_name FunBox}, @{const_name PairBox},
                  @{const_name Quot}, @{const_name Zero_Rep},
                  @{const_name Suc_Rep}] s orelse
-  let val (x as (s, T)) = (s, unarize_and_unbox_type T) in
+  let val (x as (_, T)) = (s, unarize_and_unbox_type T) in
     Refute.is_IDT_constructor thy x orelse is_record_constr x orelse
     (is_abs_fun thy x andalso is_pure_typedef thy (range_type T)) orelse
     is_coconstr thy x
@@ -747,7 +751,7 @@
             (map (box_type hol_ctxt InPair) Ts))
   | _ => false
 (* hol_context -> boxability -> string * typ list -> string *)
-and should_box_type (hol_ctxt as {thy, boxes, ...}) boxy (z as (s, Ts)) =
+and should_box_type (hol_ctxt as {thy, boxes, ...}) boxy z =
   case triple_lookup (type_match thy) boxes (Type z) of
     SOME (SOME box_me) => box_me
   | _ => is_boxing_worth_it hol_ctxt boxy (Type z)
@@ -934,9 +938,9 @@
       Abs (Name.uu, dataT, @{const True})
   end
 (* hol_context -> styp -> term -> term *)
-fun discriminate_value (hol_ctxt as {thy, ...}) (x as (_, T)) t =
-  case strip_comb t of
-    (Const x', args) =>
+fun discriminate_value (hol_ctxt as {thy, ...}) x t =
+  case head_of t of
+    Const x' =>
     if x = x' then @{const True}
     else if is_constr_like thy x' then @{const False}
     else betapply (discr_term_for_constr hol_ctxt x, t)
@@ -979,7 +983,7 @@
   | construct_value thy stds (x as (s, _)) args =
     let val args = map Envir.eta_contract args in
       case hd args of
-        Const (x' as (s', _)) $ t =>
+        Const (s', _) $ t =>
         if is_sel_like_and_no_discr s' andalso
            constr_name_for_sel_like s' = s andalso
            forall (fn (n, t') =>
@@ -1063,9 +1067,8 @@
           | constrs =>
             let
               val constr_cards =
-                datatype_constrs hol_ctxt T
-                |> map (Integer.prod o map (aux (T :: avoid)) o binder_types
-                        o snd)
+                map (Integer.prod o map (aux (T :: avoid)) o binder_types o snd)
+                    constrs
             in
               if exists (curry (op =) 0) constr_cards then 0
               else Integer.sum constr_cards
@@ -1140,6 +1143,8 @@
 fun is_built_in_theory thy = Theory.subthy (thy, @{theory Refute})
 
 (* term -> bool *)
+val is_trivial_definition =
+  the_default false o try (op aconv o Logic.dest_equals)
 val is_plain_definition =
   let
     (* term -> bool *)
@@ -1177,7 +1182,9 @@
     val defs =
       (thy |> PureThy.all_thms_of
            |> filter (curry (op =) Thm.definitionK o Thm.get_kind o snd)
-           |> map (prop_of o snd) |> filter is_plain_definition) @
+           |> map (prop_of o snd)
+           |> filter_out is_trivial_definition
+           |> filter is_plain_definition) @
       user_defs @ built_in_defs
   in (defs, built_in_nondefs, user_nondefs) end
 
@@ -1199,10 +1206,15 @@
                  (s, unarize_type T) of
           SOME n => SOME n
         | NONE =>
-          if is_fun_type T andalso is_set_type (domain_type T) then
-            AList.lookup (op =) built_in_set_consts s
-          else
-            NONE
+          case s of
+            @{const_name zero_class.zero} =>
+            if is_iterator_type T then SOME 0 else NONE
+          | @{const_name Suc} =>
+            if is_iterator_type (domain_type T) then SOME 0 else NONE
+          | _ => if is_fun_type T andalso is_set_type (domain_type T) then
+                   AList.lookup (op =) built_in_set_consts s
+                 else
+                   NONE
     end
 (* theory -> (typ option * bool) list -> bool -> styp -> bool *)
 val is_built_in_const = is_some oooo arity_of_built_in_const
@@ -1233,12 +1245,12 @@
     |> map_filter (try (Refute.specialize_type thy x))
     |> filter (curry (op =) (Const x) o term_under_def)
 
-(* theory -> term -> term option *)
-fun normalized_rhs_of thy t =
+(* term -> term option *)
+fun normalized_rhs_of t =
   let
     (* term option -> term option *)
     fun aux (v as Var _) (SOME t) = SOME (lambda v t)
-      | aux (c as Const (@{const_name TYPE}, T)) (SOME t) = SOME (lambda c t)
+      | aux (c as Const (@{const_name TYPE}, _)) (SOME t) = SOME (lambda c t)
       | aux _ _ = NONE
     val (lhs, rhs) =
       case t of
@@ -1256,7 +1268,7 @@
     NONE
   else
     x |> def_props_for_const thy [(NONE, false)] false table |> List.last
-      |> normalized_rhs_of thy |> Option.map (prefix_abs_vars s)
+      |> normalized_rhs_of |> Option.map (prefix_abs_vars s)
     handle List.Empty => NONE
 
 (* term -> fixpoint_kind *)
@@ -1366,13 +1378,12 @@
     ||> single |> op ::
   end
 (* theory -> string * typ list -> term list *)
-fun optimized_typedef_axioms thy (abs_z as (abs_s, abs_Ts)) =
+fun optimized_typedef_axioms thy (abs_z as (abs_s, _)) =
   let val abs_T = Type abs_z in
     if is_univ_typedef thy abs_T then
       []
     else case typedef_info thy abs_s of
-      SOME {abs_type, rep_type, Abs_name, Rep_name, prop_of_Rep, set_name,
-            ...} =>
+      SOME {abs_type, rep_type, Rep_name, prop_of_Rep, set_name, ...} =>
       let
         val rep_T = instantiate_type thy abs_type abs_T rep_type
         val rep_t = Const (Rep_name, abs_T --> rep_T)
@@ -1495,7 +1506,7 @@
          [!simp_table, psimp_table]
 fun is_inductive_pred hol_ctxt =
   is_real_inductive_pred hol_ctxt andf (not o is_real_equational_fun hol_ctxt)
-fun is_equational_fun (hol_ctxt as {thy, def_table, ...}) =
+fun is_equational_fun hol_ctxt =
   (is_real_equational_fun hol_ctxt orf is_real_inductive_pred hol_ctxt
    orf (String.isPrefix ubfp_prefix orf String.isPrefix lbfp_prefix) o fst)
 
@@ -1522,7 +1533,7 @@
   | is_constr_pattern _ (Var _) = true
   | is_constr_pattern thy t =
     case strip_comb t of
-      (Const (x as (s, _)), args) =>
+      (Const x, args) =>
       is_constr_like thy x andalso forall (is_constr_pattern thy) args
     | _ => false
 fun is_constr_pattern_lhs thy t =
@@ -1536,9 +1547,9 @@
 val unfold_max_depth = 255
 
 (* hol_context -> term -> term *)
-fun unfold_defs_in_term (hol_ctxt as {thy, stds, destroy_constrs, fast_descrs,
-                                      case_names, def_table, ground_thm_table,
-                                      ersatz_table, ...}) =
+fun unfold_defs_in_term (hol_ctxt as {thy, ctxt, stds, fast_descrs, case_names,
+                                      def_table, ground_thm_table, ersatz_table,
+                                      ...}) =
   let
     (* int -> typ list -> term -> term *)
     fun do_term depth Ts t =
@@ -1566,8 +1577,7 @@
          handle SAME () => betapply (do_term depth Ts t0, do_term depth Ts t1))
       | Const (@{const_name refl_on}, T) $ Const (@{const_name top}, _) $ t2 =>
         do_const depth Ts t (@{const_name refl'}, range_type T) [t2]
-      | (t0 as Const (x as (@{const_name Sigma}, T))) $ t1
-        $ (t2 as Abs (_, _, t2')) =>
+      | (t0 as Const (@{const_name Sigma}, _)) $ t1 $ (t2 as Abs (_, _, t2')) =>
         betapplys (t0 |> loose_bvar1 (t2', 0) ? do_term depth Ts,
                    map (do_term depth Ts) [t1, t2])
       | Const (x as (@{const_name distinct},
@@ -1575,7 +1585,7 @@
         $ (t1 as _ $ _) =>
         (t1 |> HOLogic.dest_list |> distinctness_formula T'
          handle TERM _ => do_const depth Ts t x [t1])
-      | (t0 as Const (x as (@{const_name If}, _))) $ t1 $ t2 $ t3 =>
+      | Const (x as (@{const_name If}, _)) $ t1 $ t2 $ t3 =>
         if is_ground_term t1 andalso
            exists (Pattern.matches thy o rpair t1)
                   (Inttab.lookup_list ground_thm_table (hash_term t1)) then
@@ -1622,7 +1632,7 @@
               else if is_stale_constr thy x then
                 raise NOT_SUPPORTED ("(non-co)constructors of codatatypes \
                                      \(\"" ^ s ^ "\")")
-              else if is_quot_abs_fun thy x then
+              else if is_quot_abs_fun ctxt x then
                 let
                   val rep_T = domain_type T
                   val abs_T = range_type T
@@ -1632,7 +1642,7 @@
                                $ (Const (@{const_name quot_normal},
                                          rep_T --> rep_T) $ Bound 0)), ts)
                 end
-              else if is_quot_rep_fun thy x then
+              else if is_quot_rep_fun ctxt x then
                 let
                   val abs_T = domain_type T
                   val rep_T = range_type T
@@ -1831,9 +1841,9 @@
              Unsynchronized.change wf_cache (cons (x, (gfp, wf))); wf
            end
 
-(* typ list -> typ -> typ -> term -> term *)
-fun ap_curry [_] _ _ t = t
-  | ap_curry arg_Ts tuple_T body_T t =
+(* typ list -> typ -> term -> term *)
+fun ap_curry [_] _ t = t
+  | ap_curry arg_Ts tuple_T t =
     let val n = length arg_Ts in
       list_abs (map (pair "c") arg_Ts,
                 incr_boundvars n t
@@ -1843,7 +1853,7 @@
 (* int -> term -> int *)
 fun num_occs_of_bound_in_term j (t1 $ t2) =
     op + (pairself (num_occs_of_bound_in_term j) (t1, t2))
-  | num_occs_of_bound_in_term j (Abs (s, T, t')) =
+  | num_occs_of_bound_in_term j (Abs (_, _, t')) =
     num_occs_of_bound_in_term (j + 1) t'
   | num_occs_of_bound_in_term j (Bound j') = if j' = j then 1 else 0
   | num_occs_of_bound_in_term _ _ = 0
@@ -1917,8 +1927,7 @@
   in aux end
 
 (* hol_context -> styp -> term -> term *)
-fun starred_linear_pred_const (hol_ctxt as {simp_table, ...}) (x as (s, T))
-                              def =
+fun starred_linear_pred_const (hol_ctxt as {simp_table, ...}) (s, T) def =
   let
     val j = maxidx_of_term def + 1
     val (outer, fp_app) = strip_abs def
@@ -1947,7 +1956,7 @@
                  $ (Const (@{const_name split}, curried_T --> uncurried_T)
                     $ list_comb (Const step_x, outer_bounds)))
               $ list_comb (Const base_x, outer_bounds)
-              |> ap_curry tuple_arg_Ts tuple_T bool_T)
+              |> ap_curry tuple_arg_Ts tuple_T)
     |> unfold_defs_in_term hol_ctxt
   end
 
@@ -2017,7 +2026,7 @@
 
 (* hol_context -> styp -> term list *)
 fun raw_equational_fun_axioms (hol_ctxt as {thy, stds, fast_descrs, simp_table,
-                                            psimp_table, ...}) (x as (s, _)) =
+                                            psimp_table, ...}) x =
   case def_props_for_const thy stds fast_descrs (!simp_table) x of
     [] => (case def_props_for_const thy stds fast_descrs psimp_table x of
              [] => [inductive_pred_axiom hol_ctxt x]
@@ -2044,7 +2053,7 @@
       | add_type _ table = table
     val table = fold (fold_types (fold_atyps add_type)) ts []
     (* typ -> typ *)
-    fun coalesce (TFree (s, S)) = TFree (AList.lookup (op =) table S |> the, S)
+    fun coalesce (TFree (_, S)) = TFree (AList.lookup (op =) table S |> the, S)
       | coalesce T = T
   in map (map_types (map_atyps coalesce)) ts end
 
@@ -2122,7 +2131,7 @@
 (* int list -> int list -> typ -> typ *)
 fun format_type default_format format T =
   let
-    val T = unarize_and_unbox_type T
+    val T = unarize_unbox_and_uniterize_type T
     val format = format |> filter (curry (op <) 0)
   in
     if forall (curry (op =) 1) format then
@@ -2172,14 +2181,14 @@
              ((arg_var_prefix ^ nat_subscript (n + 1), 0), nth missing_Ts n)
            val missing_vars = map nth_missing_var (0 upto length missing_js - 1)
            val vars = special_bounds ts @ missing_vars
-           val ts' = map2 (fn T => fn j =>
-                              case AList.lookup (op =) (js ~~ ts) j of
-                                SOME t => do_term t
-                              | NONE =>
-                                Var (nth missing_vars
-                                         (find_index (curry (op =) j)
-                                                     missing_js)))
-                          Ts (0 upto max_j)
+           val ts' =
+             map (fn j =>
+                     case AList.lookup (op =) (js ~~ ts) j of
+                       SOME t => do_term t
+                     | NONE =>
+                       Var (nth missing_vars
+                                (find_index (curry (op =) j) missing_js)))
+                 (0 upto max_j)
            val t = do_const x' |> fst
            val format =
              case AList.lookup (fn (SOME t1, SOME t2) => term_match thy (t1, t2)
@@ -2251,7 +2260,7 @@
          let val t = Const (original_name s, T) in
            (t, format_term_type thy def_table formats t)
          end)
-      |>> map_types unarize_and_unbox_type
+      |>> map_types unarize_unbox_and_uniterize_type
       |>> shorten_names_in_term |>> Term.map_abs_vars shortest_name
   in do_const end
 
--- a/src/HOL/Tools/Nitpick/nitpick_isar.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_isar.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -169,7 +169,7 @@
    | "false" => SOME false
    | "true" => SOME true
    | "" => SOME true
-   | s => raise Option)
+   | _ => raise Option)
   handle Option.Option =>
          let val ss = map quote ((option ? cons "smart") ["true", "false"]) in
            error ("Parameter " ^ quote name ^ " must be assigned " ^
@@ -433,7 +433,6 @@
   let
     val thy = Proof.theory_of state
     val ctxt = Proof.context_of state
-    val thm = #goal (Proof.raw_goal state)
     val _ = List.app check_raw_param override_params
     val params as {blocking, debug, ...} =
       extract_params ctxt auto (default_raw_params thy) override_params
--- a/src/HOL/Tools/Nitpick/nitpick_kodkod.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_kodkod.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -24,7 +24,7 @@
   val kk_tuple : bool -> int -> int list -> Kodkod.tuple
   val tuple_set_from_atom_schema : (int * int) list -> Kodkod.tuple_set
   val sequential_int_bounds : int -> Kodkod.int_bound list
-  val pow_of_two_int_bounds : int -> int -> int -> Kodkod.int_bound list
+  val pow_of_two_int_bounds : int -> int -> Kodkod.int_bound list
   val bounds_for_built_in_rels_in_formula :
     bool -> int -> int -> int -> int -> Kodkod.formula -> Kodkod.bound list
   val bound_for_plain_rel : Proof.context -> bool -> nut -> Kodkod.bound
@@ -36,7 +36,7 @@
     hol_context -> bool -> int -> int Typtab.table -> kodkod_constrs
     -> nut NameTable.table -> dtype_spec list -> Kodkod.formula list
   val kodkod_formula_from_nut :
-    int -> int Typtab.table -> kodkod_constrs -> nut -> Kodkod.formula
+    int Typtab.table -> kodkod_constrs -> nut -> Kodkod.formula
 end;
 
 structure Nitpick_Kodkod : NITPICK_KODKOD =
@@ -127,7 +127,7 @@
 (* int -> KK.int_bound list *)
 fun sequential_int_bounds n = [(NONE, map single_atom (index_seq 0 n))]
 (* int -> int -> KK.int_bound list *)
-fun pow_of_two_int_bounds bits j0 univ_card =
+fun pow_of_two_int_bounds bits j0 =
   let
     (* int -> int -> int -> KK.int_bound list *)
     fun aux 0  _ _ = []
@@ -141,7 +141,7 @@
 fun built_in_rels_in_formula formula =
   let
     (* KK.rel_expr -> KK.n_ary_index list -> KK.n_ary_index list *)
-    fun rel_expr_func (r as KK.Rel (x as (n, j))) =
+    fun rel_expr_func (KK.Rel (x as (n, j))) =
         if x = unsigned_bit_word_sel_rel orelse x = signed_bit_word_sel_rel then
           I
         else
@@ -304,7 +304,7 @@
         (FreeRel (x, T as Type ("fun", [T1, T2]), R as Func (Atom (_, j0), R2),
                   nick)) =
     let
-      val constr as {delta, epsilon, exclusive, explicit_max, ...} =
+      val {delta, epsilon, exclusive, explicit_max, ...} =
         constr_spec dtypes (original_name nick, T1)
     in
       ([(x, bound_comment ctxt debug nick T R)],
@@ -511,7 +511,7 @@
         raise REP ("Nitpick_Kodkod.lone_rep_fallback", [old_R, new_R])
     end
 (* kodkod_constrs -> int * int -> rep -> KK.rel_expr -> KK.rel_expr *)
-and atom_from_rel_expr kk (x as (k, j0)) old_R r =
+and atom_from_rel_expr kk x old_R r =
   case old_R of
     Func (R1, R2) =>
     let
@@ -581,7 +581,7 @@
     end
   | func_from_no_opt_rel_expr kk Unit R2 old_R r =
     (case old_R of
-       Vect (k, R') => rel_expr_from_rel_expr kk R2 R' r
+       Vect (_, R') => rel_expr_from_rel_expr kk R2 R' r
      | Func (Unit, R2') => rel_expr_from_rel_expr kk R2 R2' r
      | Func (Atom (1, _), Formula Neut) =>
        (case unopt_rep R2 of
@@ -824,8 +824,8 @@
   nfa |> graph_for_nfa |> NfaGraph.strong_conn
       |> map (fn keys => filter (member (op =) keys o fst) nfa)
 
-(* kodkod_constrs -> dtype_spec list -> nfa_table -> typ -> KK.formula *)
-fun acyclicity_axiom_for_datatype kk dtypes nfa start_T =
+(* kodkod_constrs -> nfa_table -> typ -> KK.formula *)
+fun acyclicity_axiom_for_datatype kk nfa start_T =
   #kk_no kk (#kk_intersect kk
                  (loop_path_rel_expr kk nfa (map fst nfa) start_T) KK.Iden)
 (* hol_context -> bool -> kodkod_constrs -> nut NameTable.table
@@ -834,20 +834,17 @@
   map_filter (nfa_entry_for_datatype hol_ctxt binarize kk rel_table dtypes)
              dtypes
   |> strongly_connected_sub_nfas
-  |> maps (fn nfa =>
-              map (acyclicity_axiom_for_datatype kk dtypes nfa o fst) nfa)
+  |> maps (fn nfa => map (acyclicity_axiom_for_datatype kk nfa o fst) nfa)
 
 (* hol_context -> bool -> int -> kodkod_constrs -> nut NameTable.table
    -> KK.rel_expr -> constr_spec -> int -> KK.formula *)
 fun sel_axiom_for_sel hol_ctxt binarize j0
-        (kk as {kk_all, kk_formula_if, kk_implies, kk_subset, kk_rel_eq, kk_no,
-                kk_join, ...}) rel_table dom_r
-        ({const, delta, epsilon, exclusive, explicit_max, ...} : constr_spec)
+        (kk as {kk_all, kk_formula_if, kk_subset, kk_no, kk_join, ...})
+        rel_table dom_r ({const, delta, epsilon, exclusive, ...} : constr_spec)
         n =
   let
-    val x as (_, T) =
-      binarized_and_boxed_nth_sel_for_constr hol_ctxt binarize const n
-    val (r, R, arity) = const_triple rel_table x
+    val x = binarized_and_boxed_nth_sel_for_constr hol_ctxt binarize const n
+    val (r, R, _) = const_triple rel_table x
     val R2 = dest_Func R |> snd
     val z = (epsilon - delta, delta + j0)
   in
@@ -908,10 +905,6 @@
       map (const_triple rel_table
            o binarized_and_boxed_nth_sel_for_constr hol_ctxt binarize const)
           (~1 upto num_sels - 1)
-    val j0 = case triples |> hd |> #2 of
-               Func (Atom (_, j0), _) => j0
-             | R => raise REP ("Nitpick_Kodkod.uniqueness_axiom_for_constr",
-                               [R])
     val set_r = triples |> hd |> #1
   in
     if num_sels = 0 then
@@ -962,8 +955,8 @@
   maps (other_axioms_for_datatype hol_ctxt binarize bits ofs kk rel_table)
        dtypes
 
-(* int -> int Typtab.table -> kodkod_constrs -> nut -> KK.formula *)
-fun kodkod_formula_from_nut bits ofs
+(* int Typtab.table -> kodkod_constrs -> nut -> KK.formula *)
+fun kodkod_formula_from_nut ofs
         (kk as {kk_all, kk_exist, kk_formula_let, kk_formula_if, kk_or, kk_not,
                 kk_iff, kk_implies, kk_and, kk_subset, kk_rel_eq, kk_no,
                 kk_lone, kk_one, kk_some, kk_rel_let, kk_rel_if, kk_union,
@@ -1098,7 +1091,7 @@
                 else
                   kk_subset (to_rep min_R u1) (to_rep min_R u2)
               end)
-         | Op2 (Eq, T, R, u1, u2) =>
+         | Op2 (Eq, _, _, u1, u2) =>
            (case min_rep (rep_of u1) (rep_of u2) of
               Unit => KK.True
             | Formula polar =>
@@ -1196,7 +1189,7 @@
       case u of
         Cst (False, _, Atom _) => false_atom
       | Cst (True, _, Atom _) => true_atom
-      | Cst (Iden, T, Func (Struct [R1, R2], Formula Neut)) =>
+      | Cst (Iden, _, Func (Struct [R1, R2], Formula Neut)) =>
         if R1 = R2 andalso arity_of_rep R1 = 1 then
           kk_intersect KK.Iden (kk_product (full_rel_for_rep R1) KK.Univ)
         else
@@ -1214,7 +1207,7 @@
                 (kk_rel_eq (rel_expr_from_rel_expr kk min_R R1 r1)
                            (rel_expr_from_rel_expr kk min_R R2 r2))
           end
-      | Cst (Iden, T, Func (Atom (1, j0), Formula Neut)) => KK.Atom j0
+      | Cst (Iden, _, Func (Atom (1, j0), Formula Neut)) => KK.Atom j0
       | Cst (Iden, T as Type ("fun", [T1, _]), R as Func (R1, _)) =>
         to_rep R (Cst (Iden, T, Func (one_rep ofs T1 R1, Formula Neut)))
       | Cst (Num j, T, R) =>
@@ -1235,7 +1228,7 @@
         to_bit_word_unary_op T R (curry KK.Add (KK.Num 1))
       | Cst (Suc, @{typ "nat => nat"}, Func (Atom x, _)) =>
         kk_intersect (KK.Rel suc_rel) (kk_product KK.Univ (KK.AtomSeq x))
-      | Cst (Suc, _, Func (Atom x, _)) => KK.Rel suc_rel
+      | Cst (Suc, _, Func (Atom _, _)) => KK.Rel suc_rel
       | Cst (Add, Type ("fun", [@{typ nat}, _]), _) => KK.Rel nat_add_rel
       | Cst (Add, Type ("fun", [@{typ int}, _]), _) => KK.Rel int_add_rel
       | Cst (Add, T as Type ("fun", [@{typ "unsigned_bit word"}, _]), R) =>
@@ -1303,7 +1296,7 @@
       | Cst (NatToInt, Type ("fun", [@{typ nat}, _]), Func (Atom _, Atom _)) =>
         KK.Iden
       | Cst (NatToInt, Type ("fun", [@{typ nat}, _]),
-             Func (Atom (nat_k, nat_j0), Opt (Atom (int_k, int_j0)))) =>
+             Func (Atom (_, nat_j0), Opt (Atom (int_k, int_j0)))) =>
         if nat_j0 = int_j0 then
           kk_intersect KK.Iden
               (kk_product (KK.AtomSeq (max_int_for_card int_k + 1, nat_j0))
@@ -1388,14 +1381,14 @@
         (case R of
            Func (R1, Formula Neut) => to_rep R1 u1
          | Func (Unit, Opt R) => to_guard [u1] R true_atom
-         | Func (R1, R2 as Opt _) =>
+         | Func (R1, Opt _) =>
            single_rel_rel_let kk
                (fn r => kk_rel_if (kk_no r) (empty_rel_for_rep R)
                             (rel_expr_to_func kk R1 bool_atom_R
                                               (Func (R1, Formula Neut)) r))
                (to_opt R1 u1)
          | _ => raise NUT ("Nitpick_Kodkod.to_r (SingletonSet)", [u]))
-      | Op1 (Tha, T, R, u1) =>
+      | Op1 (Tha, _, R, u1) =>
         if is_opt_rep R then
           kk_join (to_rep (Func (unopt_rep R, Opt bool_atom_R)) u1) true_atom
         else
@@ -1413,7 +1406,7 @@
       | Op2 (All, T, R as Opt _, u1, u2) =>
         to_r (Op1 (Not, T, R,
                    Op2 (Exist, T, R, u1, Op1 (Not, T, rep_of u2, u2))))
-      | Op2 (Exist, T, Opt _, u1, u2) =>
+      | Op2 (Exist, _, Opt _, u1, u2) =>
         let val rs1 = untuple to_decl u1 in
           if not (is_opt_rep (rep_of u2)) then
             kk_rel_if (kk_exist rs1 (to_f u2)) true_atom KK.None
@@ -1448,7 +1441,7 @@
                                 (int_expr_from_atom kk (type_of u1)) (r1, r2))))
                             KK.None)
                     (to_r u1) (to_r u2))
-      | Op2 (The, T, R, u1, u2) =>
+      | Op2 (The, _, R, u1, u2) =>
         if is_opt_rep R then
           let val r1 = to_opt (Func (unopt_rep R, bool_atom_R)) u1 in
             kk_rel_if (kk_one (kk_join r1 true_atom)) (kk_join r1 true_atom)
@@ -1461,7 +1454,7 @@
           let val r1 = to_rep (Func (R, Formula Neut)) u1 in
             kk_rel_if (kk_one r1) r1 (to_rep R u2)
           end
-      | Op2 (Eps, T, R, u1, u2) =>
+      | Op2 (Eps, _, R, u1, u2) =>
         if is_opt_rep (rep_of u1) then
           let
             val r1 = to_rep (Func (unopt_rep R, Opt bool_atom_R)) u1
@@ -1483,7 +1476,7 @@
                      (kk_rel_if (kk_or (kk_no r1) (kk_subset r2 r1))
                                 r2 (empty_rel_for_rep R))
           end
-      | Op2 (Triad, T, Opt (Atom (2, j0)), u1, u2) =>
+      | Op2 (Triad, _, Opt (Atom (2, j0)), u1, u2) =>
         let
           val f1 = to_f u1
           val f2 = to_f u2
@@ -1608,11 +1601,11 @@
           false_atom
         else
           to_apply R u1 u2
-      | Op2 (Lambda, T, R as Opt (Atom (1, j0)), u1, u2) =>
+      | Op2 (Lambda, _, R as Opt (Atom (1, j0)), u1, u2) =>
         to_guard [u1, u2] R (KK.Atom j0)
-      | Op2 (Lambda, T, Func (_, Formula Neut), u1, u2) =>
+      | Op2 (Lambda, _, Func (_, Formula Neut), u1, u2) =>
         kk_comprehension (untuple to_decl u1) (to_f u2)
-      | Op2 (Lambda, T, Func (_, R2), u1, u2) =>
+      | Op2 (Lambda, _, Func (_, R2), u1, u2) =>
         let
           val dom_decls = untuple to_decl u1
           val ran_schema = atom_schema_of_rep R2
@@ -1650,7 +1643,7 @@
                                (KK.Atom j0) KK.None)
          | _ => raise NUT ("Nitpick_Kodkod.to_r (Tuple)", [u]))
       | Construct ([u'], _, _, []) => to_r u'
-      | Construct (discr_u :: sel_us, T, R, arg_us) =>
+      | Construct (discr_u :: sel_us, _, _, arg_us) =>
         let
           val set_rs =
             map2 (fn sel_u => fn arg_u =>
@@ -1676,7 +1669,7 @@
         KK.DeclOne (x, KK.AtomSeq (the_single (atom_schema_of_rep R)))
       | to_decl u = raise NUT ("Nitpick_Kodkod.to_decl", [u])
     (* nut -> KK.expr_assign *)
-    and to_expr_assign (FormulaReg (j, _, R)) u =
+    and to_expr_assign (FormulaReg (j, _, _)) u =
         KK.AssignFormulaReg (j, to_f u)
       | to_expr_assign (RelReg (j, _, R)) u =
         KK.AssignRelReg ((arity_of_rep R, j), to_r u)
@@ -1775,7 +1768,7 @@
                case arity_of_rep nth_R of
                  0 => to_guard [u] res_R
                                (to_rep res_R (Cst (Unity, res_T, Unit)))
-               | arity => to_project res_R nth_R (to_rep (Opt (Struct Rs)) u) j0
+               | _ => to_project res_R nth_R (to_rep (Opt (Struct Rs)) u) j0
              end
     (* (KK.formula -> KK.formula -> KK.formula)
        -> (KK.rel_expr -> KK.rel_expr -> KK.formula) -> nut -> nut
@@ -1788,11 +1781,11 @@
       in
         case min_R of
           Vect (k, Atom _) => kk_vect_set_bool_op connective k r1 r2
-        | Func (R1, Formula Neut) => set_oper r1 r2
+        | Func (_, Formula Neut) => set_oper r1 r2
         | Func (Unit, Atom (2, j0)) =>
           connective (formula_from_atom j0 r1) (formula_from_atom j0 r2)
-        | Func (R1, Atom _) => set_oper (kk_join r1 true_atom)
-                                        (kk_join r2 true_atom)
+        | Func (_, Atom _) => set_oper (kk_join r1 true_atom)
+                                       (kk_join r2 true_atom)
         | _ => raise REP ("Nitpick_Kodkod.to_set_bool_op", [min_R])
       end
     (* (KK.formula -> KK.formula -> KK.formula)
@@ -1815,7 +1808,7 @@
              | Vect (k, Atom _) => kk_vect_set_op connective k r1 r2
              | Func (_, Formula Neut) => set_oper r1 r2
              | Func (Unit, _) => connective3 r1 r2
-             | Func (R1, _) =>
+             | Func _ =>
                double_rel_rel_let kk
                    (fn r1 => fn r2 =>
                        kk_union
@@ -1877,7 +1870,7 @@
       | Atom (1, j0) =>
         to_guard [arg_u] res_R
                  (rel_expr_from_rel_expr kk res_R (Atom (1, j0)) (to_r func_u))
-      | Atom (k, j0) =>
+      | Atom (k, _) =>
         let
           val dom_card = card_of_rep (rep_of arg_u)
           val ran_R = Atom (exact_root dom_card k,
--- a/src/HOL/Tools/Nitpick/nitpick_model.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_model.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -74,12 +74,12 @@
   |> (s <> "" andalso Symbol.is_ascii_digit (List.last (explode s)))
      ? prefix "\<^isub>,"
 (* atom_pool Unsynchronized.ref -> string -> typ -> int -> int -> string *)
-fun nth_atom_name pool prefix (T as Type (s, _)) j k =
+fun nth_atom_name pool prefix (Type (s, _)) j k =
     let val s' = shortest_name s in
       prefix ^ (if String.isPrefix "\\" s' then s' else substring (s', 0, 1)) ^
       nth_atom_suffix pool s j k
     end
-  | nth_atom_name pool prefix (T as TFree (s, _)) j k =
+  | nth_atom_name pool prefix (TFree (s, _)) j k =
     prefix ^ perhaps (try (unprefix "'")) s ^ nth_atom_suffix pool s j k
   | nth_atom_name _ _ T _ _ =
     raise TYPE ("Nitpick_Model.nth_atom_name", [T], [])
@@ -113,20 +113,23 @@
 fun unarize_and_unbox_term (Const (@{const_name FunBox}, _) $ t1) =
     unarize_and_unbox_term t1
   | unarize_and_unbox_term (Const (@{const_name PairBox},
-                                   Type ("fun", [T1, Type ("fun", [T2, T3])]))
+                                   Type ("fun", [T1, Type ("fun", [T2, _])]))
                             $ t1 $ t2) =
-    let val Ts = map unarize_and_unbox_type [T1, T2] in
+    let val Ts = map unarize_unbox_and_uniterize_type [T1, T2] in
       Const (@{const_name Pair}, Ts ---> Type ("*", Ts))
       $ unarize_and_unbox_term t1 $ unarize_and_unbox_term t2
     end
-  | unarize_and_unbox_term (Const (s, T)) = Const (s, unarize_and_unbox_type T)
+  | unarize_and_unbox_term (Const (s, T)) =
+    Const (s, unarize_unbox_and_uniterize_type T)
   | unarize_and_unbox_term (t1 $ t2) =
     unarize_and_unbox_term t1 $ unarize_and_unbox_term t2
-  | unarize_and_unbox_term (Free (s, T)) = Free (s, unarize_and_unbox_type T)
-  | unarize_and_unbox_term (Var (x, T)) = Var (x, unarize_and_unbox_type T)
+  | unarize_and_unbox_term (Free (s, T)) =
+    Free (s, unarize_unbox_and_uniterize_type T)
+  | unarize_and_unbox_term (Var (x, T)) =
+    Var (x, unarize_unbox_and_uniterize_type T)
   | unarize_and_unbox_term (Bound j) = Bound j
   | unarize_and_unbox_term (Abs (s, T, t')) =
-    Abs (s, unarize_and_unbox_type T, unarize_and_unbox_term t')
+    Abs (s, unarize_unbox_and_uniterize_type T, unarize_and_unbox_term t')
 
 (* typ -> typ -> (typ * typ) * (typ * typ) *)
 fun factor_out_types (T1 as Type ("*", [T11, T12]))
@@ -260,12 +263,12 @@
   | mk_tuple _ (t :: _) = t
   | mk_tuple T [] = raise TYPE ("Nitpick_Model.mk_tuple", [T], [])
 
-(* bool -> atom_pool -> string * string * string * string -> scope -> nut list
-   -> nut list -> nut list -> nut NameTable.table -> KK.raw_bound list -> typ
-   -> typ -> rep -> int list list -> term *)
-fun reconstruct_term unfold pool (maybe_name, base_name, step_name, abs_name)
-        ({hol_ctxt as {thy, ctxt, stds, ...}, binarize, card_assigns, bits,
-          datatypes, ofs, ...} : scope) sel_names rel_table bounds =
+(* bool -> atom_pool -> (string * string) * (string * string) -> scope
+   -> nut list -> nut list -> nut list -> nut NameTable.table
+   -> KK.raw_bound list -> typ -> typ -> rep -> int list list -> term *)
+fun reconstruct_term unfold pool ((maybe_name, abs_name), _)
+        ({hol_ctxt as {thy, stds, ...}, binarize, card_assigns, bits, datatypes,
+          ofs, ...} : scope) sel_names rel_table bounds =
   let
     val for_auto = (maybe_name = "")
     (* int list list -> int *)
@@ -357,11 +360,11 @@
       ts1 ~~ ts2 |> sort (nice_term_ord o pairself fst)
                  |> make_plain_fun maybe_opt T1 T2
                  |> unarize_and_unbox_term
-                 |> typecast_fun (unarize_and_unbox_type T')
-                                 (unarize_and_unbox_type T1)
-                                 (unarize_and_unbox_type T2)
+                 |> typecast_fun (unarize_unbox_and_uniterize_type T')
+                                 (unarize_unbox_and_uniterize_type T1)
+                                 (unarize_unbox_and_uniterize_type T2)
     (* (typ * int) list -> typ -> typ -> int -> term *)
-    fun term_for_atom seen (T as Type ("fun", [T1, T2])) T' j k =
+    fun term_for_atom seen (T as Type ("fun", [T1, T2])) T' j _ =
         let
           val k1 = card_of_type card_assigns T1
           val k2 = card_of_type card_assigns T2
@@ -525,7 +528,7 @@
                      map3 (fn T => term_for_rep seen T T) [T1, T2] [R1, R2]
                           [[js1], [js2]])
         end
-      | term_for_rep seen (Type ("fun", [T1, T2])) T' (R as Vect (k, R')) [js] =
+      | term_for_rep seen (Type ("fun", [T1, T2])) T' (Vect (k, R')) [js] =
         term_for_vect seen k R' T1 T2 T' js
       | term_for_rep seen (Type ("fun", [T1, T2])) T' (Func (R1, Formula Neut))
                      jss =
@@ -548,7 +551,7 @@
         in make_fun true T1 T2 T' ts1 ts2 end
       | term_for_rep seen T T' (Opt R) jss =
         if null jss then Const (unknown, T) else term_for_rep seen T T' R jss
-      | term_for_rep seen T _ R jss =
+      | term_for_rep _ T _ R jss =
         raise ARG ("Nitpick_Model.reconstruct_term.term_for_rep",
                    Refute.string_of_typ T ^ " " ^ string_for_rep R ^ " " ^
                    string_of_int (length jss))
@@ -559,11 +562,11 @@
 fun term_for_name pool scope sel_names rel_table bounds name =
   let val T = type_of name in
     tuple_list_for_name rel_table bounds name
-    |> reconstruct_term false pool ("", "", "", "") scope sel_names rel_table
-                        bounds T T (rep_of name)
+    |> reconstruct_term false pool (("", ""), ("", "")) scope sel_names
+                        rel_table bounds T T (rep_of name)
   end
 
-(* Proof.context -> (string * string * string * string) * Proof.context *)
+(* Proof.context -> ((string * string) * (string * string)) * Proof.context *)
 fun add_wacky_syntax ctxt =
   let
     (* term -> string *)
@@ -572,17 +575,17 @@
     val (maybe_t, thy) =
       Sign.declare_const ((@{binding nitpick_maybe}, @{typ "'a => 'a"}),
                           Mixfix (maybe_mixfix, [1000], 1000)) thy
+    val (abs_t, thy) =
+      Sign.declare_const ((@{binding nitpick_abs}, @{typ "'a => 'b"}),
+                          Mixfix (abs_mixfix, [40], 40)) thy
     val (base_t, thy) =
       Sign.declare_const ((@{binding nitpick_base}, @{typ "'a => 'a"}),
                           Mixfix (base_mixfix, [1000], 1000)) thy
     val (step_t, thy) =
       Sign.declare_const ((@{binding nitpick_step}, @{typ "'a => 'a"}),
                           Mixfix (step_mixfix, [1000], 1000)) thy
-    val (abs_t, thy) =
-      Sign.declare_const ((@{binding nitpick_abs}, @{typ "'a => 'b"}),
-                          Mixfix (abs_mixfix, [40], 40)) thy
   in
-    ((name_of maybe_t, name_of base_t, name_of step_t, name_of abs_t),
+    (pairself (pairself name_of) ((maybe_t, abs_t), (base_t, step_t)),
      ProofContext.transfer_syntax thy ctxt)
   end
 
@@ -613,18 +616,18 @@
   -> nut list -> nut list -> nut NameTable.table -> KK.raw_bound list
   -> Pretty.T * bool *)
 fun reconstruct_hol_model {show_skolems, show_datatypes, show_consts}
-        ({hol_ctxt as {thy, ctxt, max_bisim_depth, boxes, stds, wfs,
-                       user_axioms, debug, binary_ints, destroy_constrs,
-                       specialize, skolemize, star_linear_preds, uncurry,
-                       fast_descrs, tac_timeout, evals, case_names, def_table,
-                       nondef_table, user_nondefs, simp_table, psimp_table,
-                       intro_table, ground_thm_table, ersatz_table, skolems,
-                       special_funs, unrolled_preds, wf_cache, constr_cache},
+        ({hol_ctxt = {thy, ctxt, max_bisim_depth, boxes, stds, wfs, user_axioms,
+                      debug, binary_ints, destroy_constrs, specialize,
+                      skolemize, star_linear_preds, uncurry, fast_descrs,
+                      tac_timeout, evals, case_names, def_table, nondef_table,
+                      user_nondefs, simp_table, psimp_table, intro_table,
+                      ground_thm_table, ersatz_table, skolems, special_funs,
+                      unrolled_preds, wf_cache, constr_cache},
          binarize, card_assigns, bits, bisim_depth, datatypes, ofs} : scope)
         formats all_frees free_names sel_names nonsel_names rel_table bounds =
   let
     val pool = Unsynchronized.ref []
-    val (wacky_names as (_, base_name, step_name, _), ctxt) =
+    val (wacky_names as (_, base_step_names), ctxt) =
       add_wacky_syntax ctxt
     val hol_ctxt =
       {thy = thy, ctxt = ctxt, max_bisim_depth = max_bisim_depth, boxes = boxes,
@@ -679,13 +682,12 @@
         val (oper, (t1, T'), T) =
           case name of
             FreeName (s, T, _) =>
-            let val t = Free (s, unarize_and_unbox_type T) in
+            let val t = Free (s, unarize_unbox_and_uniterize_type T) in
               ("=", (t, format_term_type thy def_table formats t), T)
             end
           | ConstName (s, T, _) =>
             (assign_operator_for_const (s, T),
-             user_friendly_const hol_ctxt (base_name, step_name) formats (s, T),
-             T)
+             user_friendly_const hol_ctxt base_step_names formats (s, T), T)
           | _ => raise NUT ("Nitpick_Model.reconstruct_hol_model.\
                             \pretty_for_assign", [name])
         val t2 = if rep_of name = Any then
@@ -701,7 +703,8 @@
     (* dtype_spec -> Pretty.T *)
     fun pretty_for_datatype ({typ, card, complete, ...} : dtype_spec) =
       Pretty.block (Pretty.breaks
-          [Syntax.pretty_typ ctxt (unarize_and_unbox_type typ), Pretty.str "=",
+          [Syntax.pretty_typ ctxt (unarize_unbox_and_uniterize_type typ),
+           Pretty.str "=",
            Pretty.enum "," "{" "}"
                (map (Syntax.pretty_term ctxt) (all_values_of_type card typ) @
                 (if complete then [] else [Pretty.str unrep]))])
@@ -746,7 +749,8 @@
     val free_names =
       map (fn x as (s, T) =>
               case filter (curry (op =) x
-                       o pairf nickname_of (unarize_and_unbox_type o type_of))
+                       o pairf nickname_of
+                               (unarize_unbox_and_uniterize_type o type_of))
                        free_names of
                 [name] => name
               | [] => FreeName (s, T, Any)
@@ -767,7 +771,7 @@
 
 (* scope -> Time.time option -> nut list -> nut list -> nut NameTable.table
    -> KK.raw_bound list -> term -> bool option *)
-fun prove_hol_model (scope as {hol_ctxt as {thy, ctxt, debug, ...},
+fun prove_hol_model (scope as {hol_ctxt = {thy, ctxt, debug, ...},
                                card_assigns, ...})
                     auto_timeout free_names sel_names rel_table bounds prop =
   let
--- a/src/HOL/Tools/Nitpick/nitpick_mono.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -45,7 +45,7 @@
 exception CTYPE of string * ctype list
 
 (* string -> unit *)
-fun print_g (s : string) = ()
+fun print_g (_ : string) = ()
 
 (* var -> string *)
 val string_for_var = signed_string_of_int
@@ -265,7 +265,7 @@
   end
 
 (* cdata -> styp -> ctype *)
-fun ctype_for_constr (cdata as {hol_ctxt as {thy, ...}, alpha_T, constr_cache,
+fun ctype_for_constr (cdata as {hol_ctxt = {thy, ...}, alpha_T, constr_cache,
                                 ...}) (x as (_, T)) =
   if could_exist_alpha_sub_ctype thy alpha_T T then
     case AList.lookup (op =) (!constr_cache) x of
@@ -339,7 +339,7 @@
      | (S Minus, S Plus) => NONE
      | (V _, V _) => SOME (lits, insert (op =) (a1, a2, Leq, []) comps)
      | _ => do_sign_atom_comp Eq [] a1 a2 accum)
-  | do_sign_atom_comp cmp xs a1 a2 (accum as (lits, comps)) =
+  | do_sign_atom_comp cmp xs a1 a2 (lits, comps) =
     SOME (lits, insert (op =) (a1, a2, cmp, xs) comps)
 
 (* comp -> var list -> ctype -> ctype -> (literal list * comp list) option
@@ -367,7 +367,7 @@
     (accum |> fold (uncurry (do_ctype_comp cmp xs)) [(C11, C21), (C12, C22)]
      handle Library.UnequalLengths =>
             raise CTYPE ("Nitpick_Mono.do_ctype_comp", [C1, C2]))
-  | do_ctype_comp cmp xs (CType _) (CType _) accum =
+  | do_ctype_comp _ _ (CType _) (CType _) accum =
     accum (* no need to compare them thanks to the cache *)
   | do_ctype_comp _ _ C1 C2 _ =
     raise CTYPE ("Nitpick_Mono.do_ctype_comp", [C1, C2])
@@ -435,13 +435,6 @@
 val add_ctype_is_right_unique = add_notin_ctype_fv Minus
 val add_ctype_is_right_total = add_notin_ctype_fv Plus
 
-(* constraint_set -> constraint_set -> constraint_set *)
-fun unite (CSet (lits1, comps1, sexps1)) (CSet (lits2, comps2, sexps2)) =
-    (case SOME lits1 |> fold do_literal lits2 of
-       NONE => (print_g "**** Unsolvable"; UnsolvableCSet)
-     | SOME lits => CSet (lits, comps1 @ comps2, sexps1 @ sexps2))
-  | unite _ _ = UnsolvableCSet
-
 (* sign -> bool *)
 fun bool_from_sign Plus = false
   | bool_from_sign Minus = true
@@ -480,10 +473,6 @@
              SOME b => (x, sign_from_bool b) :: accum
            | NONE => accum) (max_var downto 1) lits
 
-(* literal list -> sign_atom -> sign option *)
-fun lookup_sign_atom _ (S sn) = SOME sn
-  | lookup_sign_atom lit (V x) = AList.lookup (op =) lit x
-
 (* comp -> string *)
 fun string_for_comp (a1, a2, cmp, xs) =
   string_for_sign_atom a1 ^ " " ^ string_for_comp_op cmp ^
@@ -522,9 +511,6 @@
       | _ => NONE
     end
 
-(* var -> constraint_set -> bool *)
-val is_solvable = is_some oo solve
-
 type ctype_schema = ctype * constraint_set
 type ctype_context =
   {bounds: ctype list,
@@ -545,8 +531,8 @@
   handle List.Empty => initial_gamma
 
 (* cdata -> term -> accumulator -> ctype * accumulator *)
-fun consider_term (cdata as {hol_ctxt as {thy, ctxt, stds, fast_descrs,
-                                          def_table, ...},
+fun consider_term (cdata as {hol_ctxt = {thy, ctxt, stds, fast_descrs,
+                                         def_table, ...},
                              alpha_T, max_fresh, ...}) =
   let
     (* typ -> ctype *)
@@ -765,7 +751,7 @@
          | Var _ => (print_g "*** Var"; unsolvable)
          | Bound j => (nth bounds j, accum)
          | Abs (_, T, @{const False}) => (ctype_for (T --> bool_T), accum)
-         | Abs (s, T, t') =>
+         | Abs (_, T, t') =>
            ((case t' of
                t1' $ Bound 0 =>
                if not (loose_bvar1 (t1', 0)) then
@@ -806,7 +792,7 @@
   in do_term end
 
 (* cdata -> sign -> term -> accumulator -> accumulator *)
-fun consider_general_formula (cdata as {hol_ctxt as {ctxt, ...}, ...}) =
+fun consider_general_formula (cdata as {hol_ctxt = {ctxt, ...}, ...}) =
   let
     (* typ -> ctype *)
     val ctype_for = fresh_ctype_for_type cdata
@@ -814,7 +800,7 @@
     val do_term = consider_term cdata
     (* sign -> term -> accumulator -> accumulator *)
     fun do_formula _ _ (_, UnsolvableCSet) = unsolvable_accum
-      | do_formula sn t (accum as (gamma as {bounds, frees, consts}, cset)) =
+      | do_formula sn t (accum as (gamma, cset)) =
         let
           (* term -> accumulator -> accumulator *)
           val do_co_formula = do_formula sn
--- a/src/HOL/Tools/Nitpick/nitpick_nut.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -7,7 +7,6 @@
 
 signature NITPICK_NUT =
 sig
-  type special_fun = Nitpick_HOL.special_fun
   type hol_context = Nitpick_HOL.hol_context
   type scope = Nitpick_Scope.scope
   type name_pool = Nitpick_Peephole.name_pool
@@ -467,7 +466,7 @@
   | factorize z = [z]
 
 (* hol_context -> op2 -> term -> nut *)
-fun nut_from_term (hol_ctxt as {thy, stds, fast_descrs, special_funs, ...}) eq =
+fun nut_from_term (hol_ctxt as {thy, stds, fast_descrs, ...}) eq =
   let
     (* string list -> typ list -> term -> nut *)
     fun aux eq ss Ts t =
@@ -518,11 +517,21 @@
             val t1 = list_comb (t0, ts')
             val T1 = fastype_of1 (Ts, t1)
           in Op2 (Apply, range_type T1, Any, sub t1, sub t2) end
+        (* styp -> term list -> term *)
+        fun construct (x as (_, T)) ts =
+          case num_binder_types T - length ts of
+            0 => Construct (map ((fn (s', T') => ConstName (s', T', Any))
+                                  o nth_sel_for_constr x)
+                                (~1 upto num_sels_for_constr_type T - 1),
+                            body_type T, Any,
+                            ts |> map (`(curry fastype_of1 Ts))
+                               |> maps factorize |> map (sub o snd))
+          | k => sub (eta_expand Ts t k)
       in
         case strip_comb t of
           (Const (@{const_name all}, _), [Abs (s, T, t1)]) =>
           do_quantifier All s T t1
-        | (t0 as Const (@{const_name all}, T), [t1]) =>
+        | (t0 as Const (@{const_name all}, _), [t1]) =>
           sub' (t0 $ eta_expand Ts t1 1)
         | (Const (@{const_name "=="}, T), [t1, t2]) => sub_equals T t1 t2
         | (Const (@{const_name "==>"}, _), [t1, t2]) =>
@@ -538,11 +547,11 @@
         | (Const (@{const_name True}, T), []) => Cst (True, T, Any)
         | (Const (@{const_name All}, _), [Abs (s, T, t1)]) =>
           do_quantifier All s T t1
-        | (t0 as Const (@{const_name All}, T), [t1]) =>
+        | (t0 as Const (@{const_name All}, _), [t1]) =>
           sub' (t0 $ eta_expand Ts t1 1)
         | (Const (@{const_name Ex}, _), [Abs (s, T, t1)]) =>
           do_quantifier Exist s T t1
-        | (t0 as Const (@{const_name Ex}, T), [t1]) =>
+        | (t0 as Const (@{const_name Ex}, _), [t1]) =>
           sub' (t0 $ eta_expand Ts t1 1)
         | (t0 as Const (@{const_name The}, T), [t1]) =>
           if fast_descrs then
@@ -568,7 +577,7 @@
         | (Const (@{const_name Let}, T), [t1, Abs (s, T', t2)]) =>
           Op3 (Let, nth_range_type 2 T, Any, BoundName (length Ts, T', Any, s),
                sub t1, sub_abs s T' t2)
-        | (t0 as Const (@{const_name Let}, T), [t1, t2]) =>
+        | (t0 as Const (@{const_name Let}, _), [t1, t2]) =>
           sub (t0 $ t1 $ eta_expand Ts t2 1)
         | (@{const Unity}, []) => Cst (Unity, @{typ unit}, Any)
         | (Const (@{const_name Pair}, T), [t1, t2]) =>
@@ -595,7 +604,10 @@
           Op2 (Product, nth_range_type 2 T, Any, sub t1, sub_abs s T' t2')
         | (Const (@{const_name image}, T), [t1, t2]) =>
           Op2 (Image, nth_range_type 2 T, Any, sub t1, sub t2)
-        | (Const (@{const_name Suc}, T), []) => Cst (Suc, T, Any)
+        | (Const (x as (s as @{const_name Suc}, T)), []) =>
+          if is_built_in_const thy stds false x then Cst (Suc, T, Any)
+          else if is_constr thy stds x then construct x []
+          else ConstName (s, T, Any)
         | (Const (@{const_name finite}, T), [t1]) =>
           (if is_finite_type hol_ctxt (domain_type T) then
              Cst (True, bool_T, Any)
@@ -604,14 +616,9 @@
            | _ => Op1 (Finite, bool_T, Any, sub t1))
         | (Const (@{const_name nat}, T), []) => Cst (IntToNat, T, Any)
         | (Const (x as (s as @{const_name zero_class.zero}, T)), []) =>
-          if is_built_in_const thy stds false x then
-            Cst (Num 0, T, Any)
-          else if is_constr thy stds x then
-            let val (s', T') = discr_for_constr x in
-              Construct ([ConstName (s', T', Any)], T, Any, [])
-            end
-          else
-            ConstName (s, T, Any)
+          if is_built_in_const thy stds false x then Cst (Num 0, T, Any)
+          else if is_constr thy stds x then construct x []
+          else ConstName (s, T, Any)
         | (Const (x as (s as @{const_name one_class.one}, T)), []) =>
           if is_built_in_const thy stds false x then Cst (Num 1, T, Any)
           else ConstName (s, T, Any)
@@ -631,7 +638,7 @@
         | (Const (x as (s as @{const_name div_class.div}, T)), []) =>
           if is_built_in_const thy stds false x then Cst (Divide, T, Any)
           else ConstName (s, T, Any)
-        | (t0 as Const (x as (s as @{const_name ord_class.less}, T)),
+        | (t0 as Const (x as (@{const_name ord_class.less}, _)),
            ts as [t1, t2]) =>
           if is_built_in_const thy stds false x then
             Op2 (Less, bool_T, Any, sub t1, sub t2)
@@ -642,7 +649,7 @@
            [t1, t2]) =>
           Op2 (Subset, bool_T, Any, sub t1, sub t2)
         (* FIXME: find out if this case is necessary *)
-        | (t0 as Const (x as (s as @{const_name ord_class.less_eq}, T)),
+        | (t0 as Const (x as (@{const_name ord_class.less_eq}, _)),
            ts as [t1, t2]) =>
           if is_built_in_const thy stds false x then
             Op1 (Not, bool_T, Any, Op2 (Less, bool_T, Any, sub t2, sub t1))
@@ -660,7 +667,7 @@
           else
             ConstName (s, T, Any)
         | (Const (@{const_name unknown}, T), []) => Cst (Unknown, T, Any)
-        | (Const (@{const_name is_unknown}, T), [t1]) =>
+        | (Const (@{const_name is_unknown}, _), [t1]) =>
           Op1 (IsUnknown, bool_T, Any, sub t1)
         | (Const (@{const_name Tha}, Type ("fun", [_, T2])), [t1]) =>
           Op1 (Tha, T2, Any, sub t1)
@@ -681,14 +688,7 @@
           Op2 (Union, T1, Any, sub t1, sub t2)
         | (t0 as Const (x as (s, T)), ts) =>
           if is_constr thy stds x then
-            case num_binder_types T - length ts of
-              0 => Construct (map ((fn (s, T) => ConstName (s, T, Any))
-                                    o nth_sel_for_constr x)
-                                  (~1 upto num_sels_for_constr_type T - 1),
-                              body_type T, Any,
-                              ts |> map (`(curry fastype_of1 Ts))
-                                 |> maps factorize |> map (sub o snd))
-            | k => sub (eta_expand Ts t k)
+            construct x ts
           else if String.isPrefix numeral_prefix s then
             Cst (Num (the (Int.fromString (unprefix numeral_prefix s))), T, Any)
           else
@@ -726,8 +726,8 @@
   in (v :: vs, NameTable.update (v, R) table) end
 (* scope -> bool -> nut -> nut list * rep NameTable.table
    -> nut list * rep NameTable.table *)
-fun choose_rep_for_const (scope as {hol_ctxt as {thy, ctxt, ...}, datatypes,
-                                    ofs, ...}) all_exact v (vs, table) =
+fun choose_rep_for_const (scope as {hol_ctxt = {thy, ...}, ...}) all_exact v
+                         (vs, table) =
   let
     val x as (s, T) = (nickname_of v, type_of v)
     val R = (if is_abs_fun thy x then
@@ -904,8 +904,7 @@
   | untuple f u = if rep_of u = Unit then [] else [f u]
 
 (* scope -> bool -> rep NameTable.table -> bool -> nut -> nut *)
-fun choose_reps_in_nut (scope as {hol_ctxt as {thy, ctxt, ...}, card_assigns,
-                                  bits, datatypes, ofs, ...})
+fun choose_reps_in_nut (scope as {card_assigns, bits, datatypes, ofs, ...})
                        unsound table def =
   let
     val bool_atom_R = Atom (2, offset_of_type ofs bool_T)
@@ -991,7 +990,7 @@
             Cst (cst, T, best_set_rep_for_type scope T)
         | Op1 (Not, T, _, u1) =>
           (case gsub def (flip_polarity polar) u1 of
-             Op2 (Triad, T, R, u11, u12) =>
+             Op2 (Triad, T, _, u11, u12) =>
              triad (s_op1 Not T (Formula Pos) u12)
                    (s_op1 Not T (Formula Neg) u11)
            | u1' => s_op1 Not T (flip_rep_polarity (rep_of u1')) u1')
@@ -1138,7 +1137,7 @@
                  else
                    unopt_rep R
              in s_op2 Lambda T R u1' u2' end
-           | R => raise NUT ("Nitpick_Nut.aux.choose_reps_in_nut", [u]))
+           | _ => raise NUT ("Nitpick_Nut.aux.choose_reps_in_nut", [u]))
         | Op2 (oper, T, _, u1, u2) =>
           if oper = All orelse oper = Exist then
             let
@@ -1307,7 +1306,7 @@
     end
   | shape_tuple (T as Type ("fun", [_, T2])) (R as Vect (k, R')) us =
     Tuple (T, R, map (shape_tuple T2 R') (batch_list (length us div k) us))
-  | shape_tuple T R [u] =
+  | shape_tuple T _ [u] =
     if type_of u = T then u else raise NUT ("Nitpick_Nut.shape_tuple", [u])
   | shape_tuple T Unit [] = Cst (Unity, T, Unit)
   | shape_tuple _ _ us = raise NUT ("Nitpick_Nut.shape_tuple", us)
@@ -1390,7 +1389,6 @@
       let
         val bs = untuple I u1
         val (_, pool, table') = fold rename_plain_var bs ([], pool, table)
-        val u11 = rename_vars_in_nut pool table' u1
       in
         Op3 (Let, T, R, rename_vars_in_nut pool table' u1,
              rename_vars_in_nut pool table u2,
--- a/src/HOL/Tools/Nitpick/nitpick_peephole.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_peephole.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -261,14 +261,10 @@
 (* bool -> int -> int -> int -> kodkod_constrs *)
 fun kodkod_constrs optim nat_card int_card main_j0 =
   let
-    val false_atom = Atom main_j0
-    val true_atom = Atom (main_j0 + 1)
-
     (* bool -> int *)
     val from_bool = atom_for_bool main_j0
     (* int -> rel_expr *)
     fun from_nat n = Atom (n + main_j0)
-    val from_int = Atom o atom_for_int (int_card, main_j0)
     (* int -> int *)
     fun to_nat j = j - main_j0
     val to_int = int_for_atom (int_card, main_j0)
@@ -344,24 +340,23 @@
                else Project (r, is)
              end
 
+    (* (rel_expr -> formula) -> rel_expr -> formula *)
+    fun s_xone xone r =
+      if is_one_rel_expr r then
+        True
+      else case arity_of_rel_expr r of
+        1 => xone r
+      | arity => foldl1 And (map (xone o s_project r o single o Num)
+                                 (index_seq 0 arity))
     (* rel_expr -> formula *)
     fun s_no None = True
       | s_no (Product (r1, r2)) = s_or (s_no r1) (s_no r2)
       | s_no (Intersect (Closure (Rel x), Iden)) = Acyclic x
       | s_no r = if is_one_rel_expr r then False else No r
     fun s_lone None = True
-      | s_lone r = if is_one_rel_expr r then True else Lone r
+      | s_lone r = s_xone Lone r
     fun s_one None = False
-      | s_one r =
-        if is_one_rel_expr r then
-          True
-        else if inline_rel_expr r then
-          case arity_of_rel_expr r of
-            1 => One r
-          | arity => foldl1 And (map (One o s_project r o single o Num)
-                                     (index_seq 0 arity))
-        else
-          One r
+      | s_one r = s_xone One r
     fun s_some None = False
       | s_some (Atom _) = True
       | s_some (Product (r1, r2)) = s_and (s_some r1) (s_some r2)
@@ -415,7 +410,7 @@
     fun s_subset (Atom j1) (Atom j2) = formula_for_bool (j1 = j2)
       | s_subset (Atom j) (AtomSeq (k, j0)) =
         formula_for_bool (j >= j0 andalso j < j0 + k)
-      | s_subset (r1 as Union (r11, r12)) r2 =
+      | s_subset (Union (r11, r12)) r2 =
         s_and (s_subset r11 r2) (s_subset r12 r2)
       | s_subset r1 (r2 as Union (r21, r22)) =
         if is_one_rel_expr r1 then
@@ -516,7 +511,7 @@
       | s_join (r1 as RelIf (f, r11, r12)) r2 =
         if inline_rel_expr r2 then s_rel_if f (s_join r11 r2) (s_join r12 r2)
         else Join (r1, r2)
-      | s_join (r1 as Atom j1) (r2 as Rel (x as (2, j2))) =
+      | s_join (r1 as Atom j1) (r2 as Rel (x as (2, _))) =
         if x = suc_rel then
           let val n = to_nat j1 + 1 in
             if n < nat_card then from_nat n else None
@@ -528,7 +523,7 @@
           s_project (s_join r21 r1) is
         else
           Join (r1, r2)
-      | s_join r1 (Join (r21, r22 as Rel (x as (3, j22)))) =
+      | s_join r1 (Join (r21, r22 as Rel (x as (3, _)))) =
         ((if x = nat_add_rel then
             case (r21, r1) of
               (Atom j1, Atom j2) =>
@@ -613,7 +608,6 @@
       in aux (arity_of_rel_expr r) r end
 
     (* rel_expr -> rel_expr -> rel_expr *)
-    fun s_nat_subtract r1 r2 = fold s_join [r1, r2] (Rel nat_subtract_rel)
     fun s_nat_less (Atom j1) (Atom j2) = from_bool (j1 < j2)
       | s_nat_less r1 r2 = fold s_join [r1, r2] (Rel nat_less_rel)
     fun s_int_less (Atom j1) (Atom j2) = from_bool (to_int j1 < to_int j2)
@@ -624,7 +618,6 @@
     (* rel_expr -> rel_expr *)
     fun d_not3 r = Join (r, Rel not3_rel)
     (* rel_expr -> rel_expr -> rel_expr *)
-    fun d_nat_subtract r1 r2 = List.foldl Join (Rel nat_subtract_rel) [r1, r2]
     fun d_nat_less r1 r2 = List.foldl Join (Rel nat_less_rel) [r1, r2]
     fun d_int_less r1 r2 = List.foldl Join (Rel int_less_rel) [r1, r2]
   in
--- a/src/HOL/Tools/Nitpick/nitpick_preproc.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_preproc.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -87,10 +87,9 @@
            SOME n =>
            if n >= 2 then
              let
-               val (arg_Ts, rest_T) = strip_n_binders n T
+               val arg_Ts = strip_n_binders n T |> fst
                val j =
-                 if hd arg_Ts = @{typ bisim_iterator} orelse
-                    is_fp_iterator_type (hd arg_Ts) then
+                 if is_iterator_type (hd arg_Ts) then
                    1
                  else case find_index (not_equal bool_T) arg_Ts of
                    ~1 => n
@@ -294,8 +293,8 @@
         $ do_term new_Ts old_Ts polar t2
       | Const (s as @{const_name The}, T) => do_description_operator s T
       | Const (s as @{const_name Eps}, T) => do_description_operator s T
-      | Const (@{const_name quot_normal}, Type ("fun", [_, T2])) =>
-        let val T' = box_type hol_ctxt InSel T2 in
+      | Const (@{const_name quot_normal}, Type ("fun", [T, _])) =>
+        let val T' = box_type hol_ctxt InFunLHS T in
           Const (@{const_name quot_normal}, T' --> T')
         end
       | Const (s as @{const_name Tha}, T) => do_description_operator s T
@@ -363,8 +362,8 @@
 fun fresh_value_var Ts k n j t =
   Var ((val_var_prefix ^ nat_subscript (n - j), k), fastype_of1 (Ts, t))
 
-(* typ list -> int -> term -> bool *)
-fun has_heavy_bounds_or_vars Ts level t =
+(* typ list -> term -> bool *)
+fun has_heavy_bounds_or_vars Ts t =
   let
     (* typ list -> bool *)
     fun aux [] = false
@@ -381,7 +380,7 @@
       Const x =>
       if not relax andalso is_constr thy stds x andalso
          not (is_fun_type (fastype_of1 (Ts, t_comb))) andalso
-         has_heavy_bounds_or_vars Ts level t_comb andalso
+         has_heavy_bounds_or_vars Ts t_comb andalso
          not (loose_bvar (t_comb, level)) then
         let
           val (j, seen) = case find_index (curry (op =) t_comb) seen of
@@ -629,7 +628,7 @@
            $ Abs (s, T, kill ss Ts ts))
       | kill _ _ _ = raise UnequalLengths
     (* string list -> typ list -> term -> term *)
-    fun gather ss Ts ((t0 as Const (@{const_name Ex}, _)) $ Abs (s1, T1, t1)) =
+    fun gather ss Ts (Const (@{const_name Ex}, _) $ Abs (s1, T1, t1)) =
         gather (ss @ [s1]) (Ts @ [T1]) t1
       | gather [] [] (Abs (s, T, t1)) = Abs (s, T, gather [] [] t1)
       | gather [] [] (t1 $ t2) = gather [] [] t1 $ gather [] [] t2
@@ -704,7 +703,7 @@
         | @{const "op -->"} $ t1 $ t2 =>
           @{const "op -->"} $ aux ss Ts js depth (flip_polarity polar) t1
           $ aux ss Ts js depth polar t2
-        | (t0 as Const (@{const_name Let}, T0)) $ t1 $ t2 =>
+        | (t0 as Const (@{const_name Let}, _)) $ t1 $ t2 =>
           t0 $ t1 $ aux ss Ts js depth polar t2
         | Const (x as (s, T)) =>
           if is_inductive_pred hol_ctxt x andalso
@@ -843,7 +842,7 @@
 val bound_var_prefix = "b"
 
 (* hol_context -> int -> term -> term *)
-fun specialize_consts_in_term (hol_ctxt as {thy, specialize, simp_table,
+fun specialize_consts_in_term (hol_ctxt as {specialize, simp_table,
                                             special_funs, ...}) depth t =
   if not specialize orelse depth > special_max_depth then
     t
@@ -919,8 +918,8 @@
 
 val cong_var_prefix = "c"
 
-(* styp -> special_triple -> special_triple -> term *)
-fun special_congruence_axiom (s, T) (js1, ts1, x1) (js2, ts2, x2) =
+(* typ -> special_triple -> special_triple -> term *)
+fun special_congruence_axiom T (js1, ts1, x1) (js2, ts2, x2) =
   let
     val (bounds1, bounds2) = pairself (map Var o special_bounds) (ts1, ts2)
     val Ts = binder_types T
@@ -959,28 +958,28 @@
     fun is_more_specific (j1, t1, x1) (j2, t2, x2) =
       x1 <> x2 andalso OrdList.subset (prod_ord int_ord TermOrd.term_ord)
                                       (j2 ~~ t2, j1 ~~ t1)
-    (* styp -> special_triple list -> special_triple list -> special_triple list
+    (* typ -> special_triple list -> special_triple list -> special_triple list
        -> term list -> term list *)
     fun do_pass_1 _ [] [_] [_] = I
-      | do_pass_1 x skipped _ [] = do_pass_2 x skipped
-      | do_pass_1 x skipped all (z :: zs) =
+      | do_pass_1 T skipped _ [] = do_pass_2 T skipped
+      | do_pass_1 T skipped all (z :: zs) =
         case filter (is_more_specific z) all
              |> sort (int_ord o pairself generality) of
-          [] => do_pass_1 x (z :: skipped) all zs
-        | (z' :: _) => cons (special_congruence_axiom x z z')
-                       #> do_pass_1 x skipped all zs
-    (* styp -> special_triple list -> term list -> term list *)
+          [] => do_pass_1 T (z :: skipped) all zs
+        | (z' :: _) => cons (special_congruence_axiom T z z')
+                       #> do_pass_1 T skipped all zs
+    (* typ -> special_triple list -> term list -> term list *)
     and do_pass_2 _ [] = I
-      | do_pass_2 x (z :: zs) =
-        fold (cons o special_congruence_axiom x z) zs #> do_pass_2 x zs
-  in fold (fn (x, zs) => do_pass_1 x [] zs zs) groups [] end
+      | do_pass_2 T (z :: zs) =
+        fold (cons o special_congruence_axiom T z) zs #> do_pass_2 T zs
+  in fold (fn ((_, T), zs) => do_pass_1 T [] zs zs) groups [] end
 
 (** Axiom selection **)
 
 (* Similar to "Refute.specialize_type" but returns all matches rather than only
    the first (preorder) match. *)
 (* theory -> styp -> term -> term list *)
-fun multi_specialize_type thy slack (x as (s, T)) t =
+fun multi_specialize_type thy slack (s, T) t =
   let
     (* term -> (typ * term) list -> (typ * term) list *)
     fun aux (Const (s', T')) ys =
@@ -1062,7 +1061,7 @@
             is_built_in_const thy stds fast_descrs x then
            accum
          else
-           let val accum as (xs, _) = (x :: xs, axs) in
+           let val accum = (x :: xs, axs) in
              if depth > axioms_max_depth then
                raise TOO_LARGE ("Nitpick_Preproc.axioms_for_term.\
                                 \add_axioms_for_term",
@@ -1130,7 +1129,7 @@
       | @{typ unit} => I
       | TFree (_, S) => add_axioms_for_sort depth T S
       | TVar (_, S) => add_axioms_for_sort depth T S
-      | Type (z as (s, Ts)) =>
+      | Type (z as (_, Ts)) =>
         fold (add_axioms_for_type depth) Ts
         #> (if is_pure_typedef thy T then
               fold (add_maybe_def_axiom depth) (optimized_typedef_axioms thy z)
@@ -1192,7 +1191,7 @@
         ((if is_constr_like thy x then
             if length args = num_binder_types T then
               case hd args of
-                Const (x' as (_, T')) $ t' =>
+                Const (_, T') $ t' =>
                 if domain_type T' = body_type T andalso
                    forall (uncurry (is_nth_sel_on t'))
                           (index_seq 0 (length args) ~~ args) then
@@ -1276,13 +1275,13 @@
    paper). *)
 val quantifier_cluster_threshold = 7
 
-(* theory -> term -> term *)
-fun push_quantifiers_inward thy =
+(* term -> term *)
+val push_quantifiers_inward =
   let
     (* string -> string list -> typ list -> term -> term *)
     fun aux quant_s ss Ts t =
       (case t of
-         (t0 as Const (s0, _)) $ Abs (s1, T1, t1 as _ $ _) =>
+         Const (s0, _) $ Abs (s1, T1, t1 as _ $ _) =>
          if s0 = quant_s then
            aux s0 (s1 :: ss) (T1 :: Ts) t1
          else if quant_s = "" andalso
@@ -1406,8 +1405,8 @@
     val table =
       Termtab.empty |> uncurry
         ? fold (add_to_uncurry_table thy) (core_t :: def_ts @ nondef_ts)
-    (* bool -> bool -> term -> term *)
-    fun do_rest def core =
+    (* bool -> term -> term *)
+    fun do_rest def =
       binarize ? binarize_nat_and_int_in_term
       #> uncurry ? uncurry_term table
       #> box ? box_fun_and_pair_in_term hol_ctxt def
@@ -1419,13 +1418,13 @@
       #> destroy_existential_equalities hol_ctxt
       #> simplify_constrs_and_sels thy
       #> distribute_quantifiers
-      #> push_quantifiers_inward thy
+      #> push_quantifiers_inward
       #> close_form
       #> Term.map_abs_vars shortest_name
   in
-    (((map (do_rest true false) def_ts, map (do_rest false false) nondef_ts),
+    (((map (do_rest true) def_ts, map (do_rest false) nondef_ts),
       (got_all_mono_user_axioms, no_poly_user_axioms)),
-     do_rest false true core_t, binarize)
+     do_rest false core_t, binarize)
   end
 
 end;
--- a/src/HOL/Tools/Nitpick/nitpick_rep.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_rep.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -166,7 +166,7 @@
 
 (* rep -> rep list *)
 fun binder_reps (Func (R1, R2)) = R1 :: binder_reps R2
-  | binder_reps R = []
+  | binder_reps _ = []
 (* rep -> rep *)
 fun body_rep (Func (_, R2)) = body_rep R2
   | body_rep R = R
@@ -272,10 +272,10 @@
     (case (best_one_rep_for_type scope T1, best_one_rep_for_type scope T2) of
        (Unit, Unit) => Unit
      | (R1, R2) => Struct [R1, R2])
-  | best_one_rep_for_type (scope as {card_assigns, datatypes, ofs, ...}) T =
+  | best_one_rep_for_type {card_assigns, datatypes, ofs, ...} T =
     (case card_of_type card_assigns T of
        1 => if is_some (datatype_spec datatypes T) orelse
-               is_fp_iterator_type T then
+               is_iterator_type T then
               Atom (1, offset_of_type ofs T)
             else
               Unit
@@ -332,7 +332,7 @@
   | type_schema_of_rep (Type ("fun", [T1, T2])) (Func (R1, R2)) =
     type_schema_of_rep T1 R1 @ type_schema_of_rep T2 R2
   | type_schema_of_rep T (Opt R) = type_schema_of_rep T R
-  | type_schema_of_rep T R = raise REP ("Nitpick_Rep.type_schema_of_rep", [R])
+  | type_schema_of_rep _ R = raise REP ("Nitpick_Rep.type_schema_of_rep", [R])
 (* typ list -> rep list -> typ list *)
 and type_schema_of_reps Ts Rs = flat (map2 type_schema_of_rep Ts Rs)
 
--- a/src/HOL/Tools/Nitpick/nitpick_scope.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_scope.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -136,7 +136,7 @@
 (* (string -> string) -> scope
    -> string list * string list * string list * string list * string list *)
 fun quintuple_for_scope quote
-        ({hol_ctxt as {thy, ctxt, stds, ...}, card_assigns, bits, bisim_depth,
+        ({hol_ctxt = {thy, ctxt, stds, ...}, card_assigns, bits, bisim_depth,
          datatypes, ...} : scope) =
   let
     val boring_Ts = [@{typ unsigned_bit}, @{typ signed_bit},
@@ -346,7 +346,7 @@
   let
     (* (typ * int) list -> (typ * int) list -> (typ * int) list option *)
     fun aux seen [] = SOME seen
-      | aux seen ((T, 0) :: _) = NONE
+      | aux _ ((_, 0) :: _) = NONE
       | aux seen ((T, k) :: rest) =
         (if is_surely_inconsistent_scope_description hol_ctxt binarize
                 ((T, k) :: seen) rest max_assigns then
@@ -359,7 +359,7 @@
   in aux [] (rev card_assigns) end
 
 (* theory -> (typ * int) list -> typ * int -> typ * int *)
-fun repair_iterator_assign thy assigns (T as Type (s, Ts), k) =
+fun repair_iterator_assign thy assigns (T as Type (_, Ts), k) =
     (T, if T = @{typ bisim_iterator} then
           let
             val co_cards = map snd (filter (is_codatatype thy o fst) assigns)
@@ -394,15 +394,15 @@
   end
   handle Option.Option => NONE
 
-(* theory -> (typ * int) list -> dtype_spec list -> int Typtab.table *)
-fun offset_table_for_card_assigns thy assigns dtypes =
+(* (typ * int) list -> dtype_spec list -> int Typtab.table *)
+fun offset_table_for_card_assigns assigns dtypes =
   let
     (* int -> (int * int) list -> (typ * int) list -> int Typtab.table
        -> int Typtab.table *)
     fun aux next _ [] = Typtab.update_new (dummyT, next)
       | aux next reusable ((T, k) :: assigns) =
-        if k = 1 orelse is_fp_iterator_type T orelse is_integer_type T
-           orelse T = @{typ bisim_iterator} orelse is_bit_type T then
+        if k = 1 orelse is_iterator_type T orelse is_integer_type T
+           orelse is_bit_type T then
           aux next reusable assigns
         else if length (these (Option.map #constrs (datatype_spec dtypes T)))
                 > 1 then
@@ -421,12 +421,10 @@
 (* scope_desc -> bool -> int -> (int -> int) -> int -> int -> bool * styp
    -> constr_spec list -> constr_spec list *)
 fun add_constr_spec (card_assigns, max_assigns) co card sum_dom_cards
-                    num_self_recs num_non_self_recs (self_rec, x as (s, T))
+                    num_self_recs num_non_self_recs (self_rec, x as (_, T))
                     constrs =
   let
     val max = constr_max max_assigns x
-    (* int -> int *)
-    fun bound k = Int.min (card, (max >= 0 ? curry Int.min max) k)
     (* unit -> int *)
     fun next_delta () = if null constrs then 0 else #epsilon (hd constrs)
     val {delta, epsilon, exclusive, total} =
@@ -496,14 +494,6 @@
      concrete = concrete, deep = deep, constrs = constrs}
   end
 
-(* int -> int *)
-fun min_bits_for_nat_value n = if n <= 0 then 0 else IntInf.log2 n + 1
-(* scope_desc -> int *)
-fun min_bits_for_max_assigns (_, []) = 0
-  | min_bits_for_max_assigns (card_assigns, max_assigns) =
-    min_bits_for_nat_value (fold Integer.max
-        (map snd card_assigns @ map snd max_assigns) 0)
-
 (* hol_context -> bool -> int -> typ list -> scope_desc -> scope *)
 fun scope_from_descriptor (hol_ctxt as {thy, stds, ...}) binarize sym_break
                           deep_dataTs (desc as (card_assigns, _)) =
@@ -520,7 +510,7 @@
     {hol_ctxt = hol_ctxt, binarize = binarize, card_assigns = card_assigns,
      datatypes = datatypes, bits = bits, bisim_depth = bisim_depth,
      ofs = if sym_break <= 0 then Typtab.empty
-           else offset_table_for_card_assigns thy card_assigns datatypes}
+           else offset_table_for_card_assigns card_assigns datatypes}
   end
 
 (* theory -> typ list -> (typ option * int list) list
--- a/src/HOL/Tools/Nitpick/nitpick_tests.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_tests.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -294,7 +294,8 @@
 *)
   ]
 
-fun problem_for_nut ctxt name u =
+(* Proof.context -> string * nut -> Kodkod.problem *)
+fun problem_for_nut ctxt (name, u) =
   let
     val debug = false
     val peephole_optim = true
@@ -308,7 +309,7 @@
                        NameTable.empty
     val u = Op1 (Not, type_of u, rep_of u, u)
             |> rename_vars_in_nut pool table
-    val formula = kodkod_formula_from_nut bits Typtab.empty constrs u
+    val formula = kodkod_formula_from_nut Typtab.empty constrs u
     val bounds = map (bound_for_plain_rel ctxt debug) free_rels
     val univ_card = univ_card nat_card int_card j0 bounds formula
     val declarative_axioms = map (declarative_axiom_for_plain_rel constrs)
@@ -320,15 +321,11 @@
      formula = formula}
   end
 
-(* string -> unit *)
-fun run_test name =
+(* unit -> unit *)
+fun run_all_tests () =
   case Kodkod.solve_any_problem false NONE 0 1
-           [problem_for_nut @{context} name
-                            (the (AList.lookup (op =) tests name))] of
+                                (map (problem_for_nut @{context}) tests) of
     Kodkod.Normal ([], _) => ()
-  | _ => warning ("Test " ^ quote name ^ " failed")
-
-(* unit -> unit *)
-fun run_all_tests () = List.app run_test (map fst tests)
+  | _ => error "Tests failed"
 
 end;
--- a/src/HOL/Tools/Nitpick/nitpick_util.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_util.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -92,7 +92,7 @@
 val max_exponent = 16384
 
 (* int -> int -> int *)
-fun reasonable_power a 0 = 1
+fun reasonable_power _ 0 = 1
   | reasonable_power a 1 = a
   | reasonable_power 0 _ = 0
   | reasonable_power 1 _ = 1
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -2436,8 +2436,8 @@
       val [polarity, depth] = additional_arguments
       val polarity' = (case prem of Prem _ => I | Negprem _ => HOLogic.mk_not | _ => I) polarity
       val depth' =
-        Const ("Algebras.minus", @{typ "code_numeral => code_numeral => code_numeral"})
-          $ depth $ Const ("Algebras.one", @{typ "Code_Numeral.code_numeral"})
+        Const (@{const_name Groups.minus}, @{typ "code_numeral => code_numeral => code_numeral"})
+          $ depth $ Const (@{const_name Groups.one}, @{typ "Code_Numeral.code_numeral"})
     in [polarity', depth'] end
   }
 
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -224,8 +224,8 @@
   @{const_name Suc}, @{const_name Nat.zero_nat_inst.zero_nat},
   @{const_name Nat.one_nat_inst.one_nat},
   @{const_name Orderings.less}, @{const_name Orderings.less_eq},
-  @{const_name Algebras.zero},
-  @{const_name Algebras.one},  @{const_name Algebras.plus},
+  @{const_name Groups.zero},
+  @{const_name Groups.one},  @{const_name Groups.plus},
   @{const_name Nat.ord_nat_inst.less_eq_nat},
   @{const_name Nat.ord_nat_inst.less_nat},
   @{const_name number_nat_inst.number_of_nat},
--- a/src/HOL/Tools/Qelim/cooper.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Qelim/cooper.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -79,9 +79,9 @@
 | Const (@{const_name Orderings.less_eq}, _) $ y $ z =>
    if term_of x aconv y then Le (Thm.dest_arg ct)
    else if term_of x aconv z then Ge (Thm.dest_arg1 ct) else Nox
-| Const (@{const_name Rings.dvd},_)$_$(Const(@{const_name Algebras.plus},_)$y$_) =>
+| Const (@{const_name Rings.dvd},_)$_$(Const(@{const_name Groups.plus},_)$y$_) =>
    if term_of x aconv y then Dvd (Thm.dest_binop ct ||> Thm.dest_arg) else Nox
-| Const (@{const_name Not},_) $ (Const (@{const_name Rings.dvd},_)$_$(Const(@{const_name Algebras.plus},_)$y$_)) =>
+| Const (@{const_name Not},_) $ (Const (@{const_name Rings.dvd},_)$_$(Const(@{const_name Groups.plus},_)$y$_)) =>
    if term_of x aconv y then
    NDvd (Thm.dest_binop (Thm.dest_arg ct) ||> Thm.dest_arg) else Nox
 | _ => Nox)
@@ -175,18 +175,18 @@
   (fn _ => EVERY [simp_tac lin_ss 1, TRY (Lin_Arith.tac ctxt 1)]);
 fun linear_cmul 0 tm = zero
   | linear_cmul n tm = case tm of
-      Const (@{const_name Algebras.plus}, _) $ a $ b => addC $ linear_cmul n a $ linear_cmul n b
-    | Const (@{const_name Algebras.times}, _) $ c $ x => mulC $ numeral1 (fn m => n * m) c $ x
-    | Const (@{const_name Algebras.minus}, _) $ a $ b => subC $ linear_cmul n a $ linear_cmul n b
-    | (m as Const (@{const_name Algebras.uminus}, _)) $ a => m $ linear_cmul n a
+      Const (@{const_name Groups.plus}, _) $ a $ b => addC $ linear_cmul n a $ linear_cmul n b
+    | Const (@{const_name Groups.times}, _) $ c $ x => mulC $ numeral1 (fn m => n * m) c $ x
+    | Const (@{const_name Groups.minus}, _) $ a $ b => subC $ linear_cmul n a $ linear_cmul n b
+    | (m as Const (@{const_name Groups.uminus}, _)) $ a => m $ linear_cmul n a
     | _ => numeral1 (fn m => n * m) tm;
 fun earlier [] x y = false
   | earlier (h::t) x y =
     if h aconv y then false else if h aconv x then true else earlier t x y;
 
 fun linear_add vars tm1 tm2 = case (tm1, tm2) of
-    (Const (@{const_name Algebras.plus}, _) $ (Const (@{const_name Algebras.times}, _) $ c1 $ x1) $ r1,
-    Const (@{const_name Algebras.plus}, _) $ (Const (@{const_name Algebras.times}, _) $ c2 $ x2) $ r2) =>
+    (Const (@{const_name Groups.plus}, _) $ (Const (@{const_name Groups.times}, _) $ c1 $ x1) $ r1,
+    Const (@{const_name Groups.plus}, _) $ (Const (@{const_name Groups.times}, _) $ c2 $ x2) $ r2) =>
    if x1 = x2 then
      let val c = numeral2 Integer.add c1 c2
       in if c = zero then linear_add vars r1 r2
@@ -194,9 +194,9 @@
      end
      else if earlier vars x1 x2 then addC $ (mulC $ c1 $ x1) $ linear_add vars r1 tm2
    else addC $ (mulC $ c2 $ x2) $ linear_add vars tm1 r2
- | (Const (@{const_name Algebras.plus}, _) $ (Const (@{const_name Algebras.times}, _) $ c1 $ x1) $ r1, _) =>
+ | (Const (@{const_name Groups.plus}, _) $ (Const (@{const_name Groups.times}, _) $ c1 $ x1) $ r1, _) =>
       addC $ (mulC $ c1 $ x1) $ linear_add vars r1 tm2
- | (_, Const (@{const_name Algebras.plus}, _) $ (Const (@{const_name Algebras.times}, _) $ c2 $ x2) $ r2) =>
+ | (_, Const (@{const_name Groups.plus}, _) $ (Const (@{const_name Groups.times}, _) $ c2 $ x2) $ r2) =>
       addC $ (mulC $ c2 $ x2) $ linear_add vars tm1 r2
  | (_, _) => numeral2 Integer.add tm1 tm2;
 
@@ -205,10 +205,10 @@
 
 
 fun lint vars tm =  if is_numeral tm then tm  else case tm of
-  Const (@{const_name Algebras.uminus}, _) $ t => linear_neg (lint vars t)
-| Const (@{const_name Algebras.plus}, _) $ s $ t => linear_add vars (lint vars s) (lint vars t)
-| Const (@{const_name Algebras.minus}, _) $ s $ t => linear_sub vars (lint vars s) (lint vars t)
-| Const (@{const_name Algebras.times}, _) $ s $ t =>
+  Const (@{const_name Groups.uminus}, _) $ t => linear_neg (lint vars t)
+| Const (@{const_name Groups.plus}, _) $ s $ t => linear_add vars (lint vars s) (lint vars t)
+| Const (@{const_name Groups.minus}, _) $ s $ t => linear_sub vars (lint vars s) (lint vars t)
+| Const (@{const_name Groups.times}, _) $ s $ t =>
   let val s' = lint vars s
       val t' = lint vars t
   in if is_numeral s' then (linear_cmul (dest_numeral s') t')
@@ -270,7 +270,7 @@
       val d'' = Thm.rhs_of dth |> Thm.dest_arg1
      in
       case tt' of
-        Const(@{const_name Algebras.plus},_)$(Const(@{const_name Algebras.times},_)$c$_)$_ =>
+        Const(@{const_name Groups.plus},_)$(Const(@{const_name Groups.times},_)$c$_)$_ =>
         let val x = dest_numeral c
         in if x < 0 then Conv.fconv_rule (arg_conv (arg_conv (lint_conv ctxt vs)))
                                        (Thm.transitive dth (inst' [d'',t'] dvd_uminus'))
@@ -293,15 +293,15 @@
   val ins = insert (op = : int * int -> bool)
   fun h (acc,dacc) t =
    case (term_of t) of
-    Const(s,_)$(Const(@{const_name Algebras.times},_)$c$y)$ _ =>
+    Const(s,_)$(Const(@{const_name Groups.times},_)$c$y)$ _ =>
     if x aconv y andalso member (op =)
       ["op =", @{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
     then (ins (dest_numeral c) acc,dacc) else (acc,dacc)
-  | Const(s,_)$_$(Const(@{const_name Algebras.times},_)$c$y) =>
+  | Const(s,_)$_$(Const(@{const_name Groups.times},_)$c$y) =>
     if x aconv y andalso member (op =)
        [@{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
     then (ins (dest_numeral c) acc, dacc) else (acc,dacc)
-  | Const(@{const_name Rings.dvd},_)$_$(Const(@{const_name Algebras.plus},_)$(Const(@{const_name Algebras.times},_)$c$y)$_) =>
+  | Const(@{const_name Rings.dvd},_)$_$(Const(@{const_name Groups.plus},_)$(Const(@{const_name Groups.times},_)$c$y)$_) =>
     if x aconv y then (acc,ins (dest_numeral c) dacc) else (acc,dacc)
   | Const("op &",_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
   | Const("op |",_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
@@ -335,15 +335,15 @@
    Const("op &",_)$_$_ => binop_conv unit_conv t
   | Const("op |",_)$_$_ => binop_conv unit_conv t
   | Const (@{const_name Not},_)$_ => arg_conv unit_conv t
-  | Const(s,_)$(Const(@{const_name Algebras.times},_)$c$y)$ _ =>
+  | Const(s,_)$(Const(@{const_name Groups.times},_)$c$y)$ _ =>
     if x=y andalso member (op =)
       ["op =", @{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
     then cv (l div dest_numeral c) t else Thm.reflexive t
-  | Const(s,_)$_$(Const(@{const_name Algebras.times},_)$c$y) =>
+  | Const(s,_)$_$(Const(@{const_name Groups.times},_)$c$y) =>
     if x=y andalso member (op =)
       [@{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
     then cv (l div dest_numeral c) t else Thm.reflexive t
-  | Const(@{const_name Rings.dvd},_)$d$(r as (Const(@{const_name Algebras.plus},_)$(Const(@{const_name Algebras.times},_)$c$y)$_)) =>
+  | Const(@{const_name Rings.dvd},_)$d$(r as (Const(@{const_name Groups.plus},_)$(Const(@{const_name Groups.times},_)$c$y)$_)) =>
     if x=y then
       let
        val k = l div dest_numeral c
@@ -546,10 +546,10 @@
   | @{term "0::int"} => C 0
   | @{term "1::int"} => C 1
   | Term.Bound i => Bound i
-  | Const(@{const_name Algebras.uminus},_)$t' => Neg (i_of_term vs t')
-  | Const(@{const_name Algebras.plus},_)$t1$t2 => Add (i_of_term vs t1,i_of_term vs t2)
-  | Const(@{const_name Algebras.minus},_)$t1$t2 => Sub (i_of_term vs t1,i_of_term vs t2)
-  | Const(@{const_name Algebras.times},_)$t1$t2 =>
+  | Const(@{const_name Groups.uminus},_)$t' => Neg (i_of_term vs t')
+  | Const(@{const_name Groups.plus},_)$t1$t2 => Add (i_of_term vs t1,i_of_term vs t2)
+  | Const(@{const_name Groups.minus},_)$t1$t2 => Sub (i_of_term vs t1,i_of_term vs t2)
+  | Const(@{const_name Groups.times},_)$t1$t2 =>
      (Mul (HOLogic.dest_number t1 |> snd, i_of_term vs t2)
     handle TERM _ =>
        (Mul (HOLogic.dest_number t2 |> snd, i_of_term vs t1)
--- a/src/HOL/Tools/Qelim/presburger.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Qelim/presburger.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -52,15 +52,15 @@
 
 local
  fun isnum t = case t of 
-   Const(@{const_name Algebras.zero},_) => true
- | Const(@{const_name Algebras.one},_) => true
+   Const(@{const_name Groups.zero},_) => true
+ | Const(@{const_name Groups.one},_) => true
  | @{term "Suc"}$s => isnum s
  | @{term "nat"}$s => isnum s
  | @{term "int"}$s => isnum s
- | Const(@{const_name Algebras.uminus},_)$s => isnum s
- | Const(@{const_name Algebras.plus},_)$l$r => isnum l andalso isnum r
- | Const(@{const_name Algebras.times},_)$l$r => isnum l andalso isnum r
- | Const(@{const_name Algebras.minus},_)$l$r => isnum l andalso isnum r
+ | Const(@{const_name Groups.uminus},_)$s => isnum s
+ | Const(@{const_name Groups.plus},_)$l$r => isnum l andalso isnum r
+ | Const(@{const_name Groups.times},_)$l$r => isnum l andalso isnum r
+ | Const(@{const_name Groups.minus},_)$l$r => isnum l andalso isnum r
  | Const(@{const_name Power.power},_)$l$r => isnum l andalso isnum r
  | Const(@{const_name Divides.mod},_)$l$r => isnum l andalso isnum r
  | Const(@{const_name Divides.div},_)$l$r => isnum l andalso isnum r
--- a/src/HOL/Tools/Quotient/quotient_info.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/Quotient/quotient_info.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -59,7 +59,7 @@
   (type T = maps_info Symtab.table
    val empty = Symtab.empty
    val extend = I
-   val merge = Symtab.merge (K true))
+   fun merge data = Symtab.merge (K true) data)
 
 fun maps_defined thy s =
   Symtab.defined (MapsData.get thy) s
@@ -123,7 +123,7 @@
   (type T = quotdata_info Symtab.table
    val empty = Symtab.empty
    val extend = I
-   val merge = Symtab.merge (K true))
+   fun merge data = Symtab.merge (K true) data)
 
 fun transform_quotdata phi {qtyp, rtyp, equiv_rel, equiv_thm} =
   {qtyp = Morphism.typ phi qtyp,
@@ -193,7 +193,7 @@
 fun qconsts_lookup thy t =
   let
     val (name, qty) = dest_Const t
-    fun matches x =
+    fun matches (x: qconsts_info) =
       let
         val (name', qty') = dest_Const (#qconst x);
       in
@@ -280,7 +280,7 @@
     OuterKeyword.diag (Scan.succeed (Toplevel.keep (pp_fn o Toplevel.context_of)))
 
 val _ = map improper_command
-  [(print_mapsinfo, "print_maps", "prints out all map functions"),
+  [(print_mapsinfo, "print_quotmaps", "prints out all map functions"),
    (print_quotinfo, "print_quotients", "prints out all quotients"),
    (print_qconstinfo, "print_quotconsts", "prints out all quotient constants")]
 
--- a/src/HOL/Tools/arith_data.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/arith_data.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -75,11 +75,11 @@
 
 fun mk_number T n = HOLogic.number_of_const T $ HOLogic.mk_numeral n;
 
-val mk_plus = HOLogic.mk_binop @{const_name Algebras.plus};
+val mk_plus = HOLogic.mk_binop @{const_name Groups.plus};
 
 fun mk_minus t = 
   let val T = Term.fastype_of t
-  in Const (@{const_name Algebras.uminus}, T --> T) $ t end;
+  in Const (@{const_name Groups.uminus}, T --> T) $ t end;
 
 (*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
 fun mk_sum T []        = mk_number T 0
@@ -91,9 +91,9 @@
   | long_mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts);
 
 (*decompose additions AND subtractions as a sum*)
-fun dest_summing (pos, Const (@{const_name Algebras.plus}, _) $ t $ u, ts) =
+fun dest_summing (pos, Const (@{const_name Groups.plus}, _) $ t $ u, ts) =
         dest_summing (pos, t, dest_summing (pos, u, ts))
-  | dest_summing (pos, Const (@{const_name Algebras.minus}, _) $ t $ u, ts) =
+  | dest_summing (pos, Const (@{const_name Groups.minus}, _) $ t $ u, ts) =
         dest_summing (pos, t, dest_summing (not pos, u, ts))
   | dest_summing (pos, t, ts) =
         if pos then t::ts else mk_minus t :: ts;
--- a/src/HOL/Tools/hologic.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/hologic.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -440,9 +440,9 @@
 
 val natT = Type ("nat", []);
 
-val zero = Const ("Algebras.zero_class.zero", natT);
+val zero = Const ("Groups.zero_class.zero", natT);
 
-fun is_zero (Const ("Algebras.zero_class.zero", _)) = true
+fun is_zero (Const ("Groups.zero_class.zero", _)) = true
   | is_zero _ = false;
 
 fun mk_Suc t = Const ("Suc", natT --> natT) $ t;
@@ -458,7 +458,7 @@
       | mk n = mk_Suc (mk (n - 1));
   in if n < 0 then raise TERM ("mk_nat: negative number", []) else mk n end;
 
-fun dest_nat (Const ("Algebras.zero_class.zero", _)) = 0
+fun dest_nat (Const ("Groups.zero_class.zero", _)) = 0
   | dest_nat (Const ("Suc", _) $ t) = dest_nat t + 1
   | dest_nat t = raise TERM ("dest_nat", [t]);
 
@@ -508,12 +508,12 @@
   | add_numerals (Abs (_, _, t)) = add_numerals t
   | add_numerals _ = I;
 
-fun mk_number T 0 = Const ("Algebras.zero_class.zero", T)
-  | mk_number T 1 = Const ("Algebras.one_class.one", T)
+fun mk_number T 0 = Const ("Groups.zero_class.zero", T)
+  | mk_number T 1 = Const ("Groups.one_class.one", T)
   | mk_number T i = number_of_const T $ mk_numeral i;
 
-fun dest_number (Const ("Algebras.zero_class.zero", T)) = (T, 0)
-  | dest_number (Const ("Algebras.one_class.one", T)) = (T, 1)
+fun dest_number (Const ("Groups.zero_class.zero", T)) = (T, 0)
+  | dest_number (Const ("Groups.one_class.one", T)) = (T, 1)
   | dest_number (Const ("Int.number_class.number_of", Type ("fun", [_, T])) $ t) =
       (T, dest_numeral t)
   | dest_number t = raise TERM ("dest_number", [t]);
--- a/src/HOL/Tools/int_arith.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/int_arith.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -49,12 +49,12 @@
   make_simproc {lhss = lhss1, name = "one_to_of_int_one_simproc",
   proc = proc1, identifier = []};
 
-fun check (Const (@{const_name Algebras.one}, @{typ int})) = false
-  | check (Const (@{const_name Algebras.one}, _)) = true
+fun check (Const (@{const_name Groups.one}, @{typ int})) = false
+  | check (Const (@{const_name Groups.one}, _)) = true
   | check (Const (s, _)) = member (op =) [@{const_name "op ="},
-      @{const_name Algebras.times}, @{const_name Algebras.uminus},
-      @{const_name Algebras.minus}, @{const_name Algebras.plus},
-      @{const_name Algebras.zero},
+      @{const_name Groups.times}, @{const_name Groups.uminus},
+      @{const_name Groups.minus}, @{const_name Groups.plus},
+      @{const_name Groups.zero},
       @{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
   | check (a $ b) = check a andalso check b
   | check _ = false;
--- a/src/HOL/Tools/lin_arith.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/lin_arith.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -138,8 +138,8 @@
 *)
 fun demult (inj_consts : (string * typ) list) : term * Rat.rat -> term option * Rat.rat =
 let
-  fun demult ((mC as Const (@{const_name Algebras.times}, _)) $ s $ t, m) =
-      (case s of Const (@{const_name Algebras.times}, _) $ s1 $ s2 =>
+  fun demult ((mC as Const (@{const_name Groups.times}, _)) $ s $ t, m) =
+      (case s of Const (@{const_name Groups.times}, _) $ s1 $ s2 =>
         (* bracketing to the right: '(s1 * s2) * t' becomes 's1 * (s2 * t)' *)
         demult (mC $ s1 $ (mC $ s2 $ t), m)
       | _ =>
@@ -170,9 +170,9 @@
         (SOME _, _) => (SOME (mC $ s $ t), m)
       | (NONE,  m') => apsnd (Rat.mult (Rat.inv m')) (demult (s, m)))
     (* terms that evaluate to numeric constants *)
-    | demult (Const (@{const_name Algebras.uminus}, _) $ t, m) = demult (t, Rat.neg m)
-    | demult (Const (@{const_name Algebras.zero}, _), m) = (NONE, Rat.zero)
-    | demult (Const (@{const_name Algebras.one}, _), m) = (NONE, m)
+    | demult (Const (@{const_name Groups.uminus}, _) $ t, m) = demult (t, Rat.neg m)
+    | demult (Const (@{const_name Groups.zero}, _), m) = (NONE, Rat.zero)
+    | demult (Const (@{const_name Groups.one}, _), m) = (NONE, m)
     (*Warning: in rare cases number_of encloses a non-numeral,
       in which case dest_numeral raises TERM; hence all the handles below.
       Same for Suc-terms that turn out not to be numerals -
@@ -196,19 +196,19 @@
   (* Turns a term 'all' and associated multiplicity 'm' into a list 'p' of
      summands and associated multiplicities, plus a constant 'i' (with implicit
      multiplicity 1) *)
-  fun poly (Const (@{const_name Algebras.plus}, _) $ s $ t,
+  fun poly (Const (@{const_name Groups.plus}, _) $ s $ t,
         m : Rat.rat, pi : (term * Rat.rat) list * Rat.rat) = poly (s, m, poly (t, m, pi))
-    | poly (all as Const (@{const_name Algebras.minus}, T) $ s $ t, m, pi) =
+    | poly (all as Const (@{const_name Groups.minus}, T) $ s $ t, m, pi) =
         if nT T then add_atom all m pi else poly (s, m, poly (t, Rat.neg m, pi))
-    | poly (all as Const (@{const_name Algebras.uminus}, T) $ t, m, pi) =
+    | poly (all as Const (@{const_name Groups.uminus}, T) $ t, m, pi) =
         if nT T then add_atom all m pi else poly (t, Rat.neg m, pi)
-    | poly (Const (@{const_name Algebras.zero}, _), _, pi) =
+    | poly (Const (@{const_name Groups.zero}, _), _, pi) =
         pi
-    | poly (Const (@{const_name Algebras.one}, _), m, (p, i)) =
+    | poly (Const (@{const_name Groups.one}, _), m, (p, i)) =
         (p, Rat.add i m)
     | poly (Const (@{const_name Suc}, _) $ t, m, (p, i)) =
         poly (t, m, (p, Rat.add i m))
-    | poly (all as Const (@{const_name Algebras.times}, _) $ _ $ _, m, pi as (p, i)) =
+    | poly (all as Const (@{const_name Groups.times}, _) $ _ $ _, m, pi as (p, i)) =
         (case demult inj_consts (all, m) of
            (NONE,   m') => (p, Rat.add i m')
          | (SOME u, m') => add_atom u m' pi)
@@ -293,7 +293,7 @@
       Const (a, _) => member (op =) [@{const_name Orderings.max},
                                     @{const_name Orderings.min},
                                     @{const_name Groups.abs},
-                                    @{const_name Algebras.minus},
+                                    @{const_name Groups.minus},
                                     "Int.nat" (*DYNAMIC BINDING!*),
                                     "Divides.div_class.mod" (*DYNAMIC BINDING!*),
                                     "Divides.div_class.div" (*DYNAMIC BINDING!*)] a
@@ -401,9 +401,9 @@
       let
         val rev_terms   = rev terms
         val terms1      = map (subst_term [(split_term, t1)]) rev_terms
-        val terms2      = map (subst_term [(split_term, Const (@{const_name Algebras.uminus},
+        val terms2      = map (subst_term [(split_term, Const (@{const_name Groups.uminus},
                             split_type --> split_type) $ t1)]) rev_terms
-        val zero        = Const (@{const_name Algebras.zero}, split_type)
+        val zero        = Const (@{const_name Groups.zero}, split_type)
         val zero_leq_t1 = Const (@{const_name Orderings.less_eq},
                             split_type --> split_type --> HOLogic.boolT) $ zero $ t1
         val t1_lt_zero  = Const (@{const_name Orderings.less},
@@ -415,12 +415,12 @@
         SOME [(Ts, subgoal1), (Ts, subgoal2)]
       end
     (* ?P (?a - ?b) = ((?a < ?b --> ?P 0) & (ALL d. ?a = ?b + d --> ?P d)) *)
-    | (Const (@{const_name Algebras.minus}, _), [t1, t2]) =>
+    | (Const (@{const_name Groups.minus}, _), [t1, t2]) =>
       let
         (* "d" in the above theorem becomes a new bound variable after NNF   *)
         (* transformation, therefore some adjustment of indices is necessary *)
         val rev_terms       = rev terms
-        val zero            = Const (@{const_name Algebras.zero}, split_type)
+        val zero            = Const (@{const_name Groups.zero}, split_type)
         val d               = Bound 0
         val terms1          = map (subst_term [(split_term, zero)]) rev_terms
         val terms2          = map (subst_term [(incr_boundvars 1 split_term, d)])
@@ -430,7 +430,7 @@
         val t1_lt_t2        = Const (@{const_name Orderings.less},
                                 split_type --> split_type --> HOLogic.boolT) $ t1 $ t2
         val t1_eq_t2_plus_d = Const ("op =", split_type --> split_type --> HOLogic.boolT) $ t1' $
-                                (Const (@{const_name Algebras.plus},
+                                (Const (@{const_name Groups.plus},
                                   split_type --> split_type --> split_type) $ t2' $ d)
         val not_false       = HOLogic.mk_Trueprop (HOLogic.Not $ HOLogic.false_const)
         val subgoal1        = (HOLogic.mk_Trueprop t1_lt_t2) :: terms1 @ [not_false]
@@ -442,8 +442,8 @@
     | (Const ("Int.nat", _), [t1]) =>
       let
         val rev_terms   = rev terms
-        val zero_int    = Const (@{const_name Algebras.zero}, HOLogic.intT)
-        val zero_nat    = Const (@{const_name Algebras.zero}, HOLogic.natT)
+        val zero_int    = Const (@{const_name Groups.zero}, HOLogic.intT)
+        val zero_nat    = Const (@{const_name Groups.zero}, HOLogic.natT)
         val n           = Bound 0
         val terms1      = map (subst_term [(incr_boundvars 1 split_term, n)])
                             (map (incr_boundvars 1) rev_terms)
@@ -465,7 +465,7 @@
     | (Const ("Divides.div_class.mod", Type ("fun", [Type ("nat", []), _])), [t1, t2]) =>
       let
         val rev_terms               = rev terms
-        val zero                    = Const (@{const_name Algebras.zero}, split_type)
+        val zero                    = Const (@{const_name Groups.zero}, split_type)
         val i                       = Bound 1
         val j                       = Bound 0
         val terms1                  = map (subst_term [(split_term, t1)]) rev_terms
@@ -480,8 +480,8 @@
         val j_lt_t2                 = Const (@{const_name Orderings.less},
                                         split_type --> split_type--> HOLogic.boolT) $ j $ t2'
         val t1_eq_t2_times_i_plus_j = Const ("op =", split_type --> split_type --> HOLogic.boolT) $ t1' $
-                                       (Const (@{const_name Algebras.plus}, split_type --> split_type --> split_type) $
-                                         (Const (@{const_name Algebras.times},
+                                       (Const (@{const_name Groups.plus}, split_type --> split_type --> split_type) $
+                                         (Const (@{const_name Groups.times},
                                            split_type --> split_type --> split_type) $ t2' $ i) $ j)
         val not_false               = HOLogic.mk_Trueprop (HOLogic.Not $ HOLogic.false_const)
         val subgoal1                = (HOLogic.mk_Trueprop t2_eq_zero) :: terms1 @ [not_false]
@@ -497,7 +497,7 @@
     | (Const ("Divides.div_class.div", Type ("fun", [Type ("nat", []), _])), [t1, t2]) =>
       let
         val rev_terms               = rev terms
-        val zero                    = Const (@{const_name Algebras.zero}, split_type)
+        val zero                    = Const (@{const_name Groups.zero}, split_type)
         val i                       = Bound 1
         val j                       = Bound 0
         val terms1                  = map (subst_term [(split_term, zero)]) rev_terms
@@ -512,8 +512,8 @@
         val j_lt_t2                 = Const (@{const_name Orderings.less},
                                         split_type --> split_type--> HOLogic.boolT) $ j $ t2'
         val t1_eq_t2_times_i_plus_j = Const ("op =", split_type --> split_type --> HOLogic.boolT) $ t1' $
-                                       (Const (@{const_name Algebras.plus}, split_type --> split_type --> split_type) $
-                                         (Const (@{const_name Algebras.times},
+                                       (Const (@{const_name Groups.plus}, split_type --> split_type --> split_type) $
+                                         (Const (@{const_name Groups.times},
                                            split_type --> split_type --> split_type) $ t2' $ i) $ j)
         val not_false               = HOLogic.mk_Trueprop (HOLogic.Not $ HOLogic.false_const)
         val subgoal1                = (HOLogic.mk_Trueprop t2_eq_zero) :: terms1 @ [not_false]
@@ -535,7 +535,7 @@
         Type ("fun", [Type ("Int.int", []), _])), [t1, t2]) =>
       let
         val rev_terms               = rev terms
-        val zero                    = Const (@{const_name Algebras.zero}, split_type)
+        val zero                    = Const (@{const_name Groups.zero}, split_type)
         val i                       = Bound 1
         val j                       = Bound 0
         val terms1                  = map (subst_term [(split_term, t1)]) rev_terms
@@ -558,8 +558,8 @@
         val t2_lt_j                 = Const (@{const_name Orderings.less},
                                         split_type --> split_type--> HOLogic.boolT) $ t2' $ j
         val t1_eq_t2_times_i_plus_j = Const ("op =", split_type --> split_type --> HOLogic.boolT) $ t1' $
-                                       (Const (@{const_name Algebras.plus}, split_type --> split_type --> split_type) $
-                                         (Const (@{const_name Algebras.times},
+                                       (Const (@{const_name Groups.plus}, split_type --> split_type --> split_type) $
+                                         (Const (@{const_name Groups.times},
                                            split_type --> split_type --> split_type) $ t2' $ i) $ j)
         val not_false               = HOLogic.mk_Trueprop (HOLogic.Not $ HOLogic.false_const)
         val subgoal1                = (HOLogic.mk_Trueprop t2_eq_zero) :: terms1 @ [not_false]
@@ -589,7 +589,7 @@
         Type ("fun", [Type ("Int.int", []), _])), [t1, t2]) =>
       let
         val rev_terms               = rev terms
-        val zero                    = Const (@{const_name Algebras.zero}, split_type)
+        val zero                    = Const (@{const_name Groups.zero}, split_type)
         val i                       = Bound 1
         val j                       = Bound 0
         val terms1                  = map (subst_term [(split_term, zero)]) rev_terms
@@ -612,8 +612,8 @@
         val t2_lt_j                 = Const (@{const_name Orderings.less},
                                         split_type --> split_type--> HOLogic.boolT) $ t2' $ j
         val t1_eq_t2_times_i_plus_j = Const ("op =", split_type --> split_type --> HOLogic.boolT) $ t1' $
-                                       (Const (@{const_name Algebras.plus}, split_type --> split_type --> split_type) $
-                                         (Const (@{const_name Algebras.times},
+                                       (Const (@{const_name Groups.plus}, split_type --> split_type --> split_type) $
+                                         (Const (@{const_name Groups.times},
                                            split_type --> split_type --> split_type) $ t2' $ i) $ j)
         val not_false               = HOLogic.mk_Trueprop (HOLogic.Not $ HOLogic.false_const)
         val subgoal1                = (HOLogic.mk_Trueprop t2_eq_zero) :: terms1 @ [not_false]
--- a/src/HOL/Tools/nat_arith.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/nat_arith.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -19,8 +19,8 @@
 
 (** abstract syntax of structure nat: 0, Suc, + **)
 
-val mk_plus = HOLogic.mk_binop @{const_name Algebras.plus};
-val dest_plus = HOLogic.dest_bin @{const_name Algebras.plus} HOLogic.natT;
+val mk_plus = HOLogic.mk_binop @{const_name Groups.plus};
+val dest_plus = HOLogic.dest_bin @{const_name Groups.plus} HOLogic.natT;
 
 fun mk_sum [] = HOLogic.zero
   | mk_sum [t] = t
@@ -85,8 +85,8 @@
 structure DiffCancelSums = CancelSumsFun
 (struct
   open CommonCancelSums;
-  val mk_bal = HOLogic.mk_binop @{const_name Algebras.minus};
-  val dest_bal = HOLogic.dest_bin @{const_name Algebras.minus} HOLogic.natT;
+  val mk_bal = HOLogic.mk_binop @{const_name Groups.minus};
+  val dest_bal = HOLogic.dest_bin @{const_name Groups.minus} HOLogic.natT;
   val uncancel_tac = gen_uncancel_tac @{thm "diff_cancel"};
 end);
 
--- a/src/HOL/Tools/nat_numeral_simprocs.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/nat_numeral_simprocs.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -32,7 +32,7 @@
   | find_first_numeral past [] = raise TERM("find_first_numeral", []);
 
 val zero = mk_number 0;
-val mk_plus = HOLogic.mk_binop @{const_name Algebras.plus};
+val mk_plus = HOLogic.mk_binop @{const_name Groups.plus};
 
 (*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
 fun mk_sum []        = zero
@@ -43,7 +43,7 @@
 fun long_mk_sum []        = HOLogic.zero
   | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
 
-val dest_plus = HOLogic.dest_bin @{const_name Algebras.plus} HOLogic.natT;
+val dest_plus = HOLogic.dest_bin @{const_name Groups.plus} HOLogic.natT;
 
 
 (** Other simproc items **)
@@ -61,14 +61,14 @@
 (*** CancelNumerals simprocs ***)
 
 val one = mk_number 1;
-val mk_times = HOLogic.mk_binop @{const_name Algebras.times};
+val mk_times = HOLogic.mk_binop @{const_name Groups.times};
 
 fun mk_prod [] = one
   | mk_prod [t] = t
   | mk_prod (t :: ts) = if t = one then mk_prod ts
                         else mk_times (t, mk_prod ts);
 
-val dest_times = HOLogic.dest_bin @{const_name Algebras.times} HOLogic.natT;
+val dest_times = HOLogic.dest_bin @{const_name Groups.times} HOLogic.natT;
 
 fun dest_prod t =
       let val (t,u) = dest_times t
@@ -194,8 +194,8 @@
 structure DiffCancelNumerals = CancelNumeralsFun
  (open CancelNumeralsCommon
   val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binop @{const_name Algebras.minus}
-  val dest_bal = HOLogic.dest_bin @{const_name Algebras.minus} HOLogic.natT
+  val mk_bal   = HOLogic.mk_binop @{const_name Groups.minus}
+  val dest_bal = HOLogic.dest_bin @{const_name Groups.minus} HOLogic.natT
   val bal_add1 = @{thm nat_diff_add_eq1} RS trans
   val bal_add2 = @{thm nat_diff_add_eq2} RS trans
 );
--- a/src/HOL/Tools/numeral_simprocs.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/numeral_simprocs.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -34,12 +34,12 @@
 val long_mk_sum = Arith_Data.long_mk_sum;
 val dest_sum = Arith_Data.dest_sum;
 
-val mk_diff = HOLogic.mk_binop @{const_name Algebras.minus};
-val dest_diff = HOLogic.dest_bin @{const_name Algebras.minus} Term.dummyT;
+val mk_diff = HOLogic.mk_binop @{const_name Groups.minus};
+val dest_diff = HOLogic.dest_bin @{const_name Groups.minus} Term.dummyT;
 
-val mk_times = HOLogic.mk_binop @{const_name Algebras.times};
+val mk_times = HOLogic.mk_binop @{const_name Groups.times};
 
-fun one_of T = Const(@{const_name Algebras.one}, T);
+fun one_of T = Const(@{const_name Groups.one}, T);
 
 (* build product with trailing 1 rather than Numeral 1 in order to avoid the
    unnecessary restriction to type class number_ring
@@ -56,7 +56,7 @@
 fun long_mk_prod T []        = one_of T
   | long_mk_prod T (t :: ts) = mk_times (t, mk_prod T ts);
 
-val dest_times = HOLogic.dest_bin @{const_name Algebras.times} Term.dummyT;
+val dest_times = HOLogic.dest_bin @{const_name Groups.times} Term.dummyT;
 
 fun dest_prod t =
       let val (t,u) = dest_times t
@@ -72,7 +72,7 @@
 fun mk_coeff (k, t) = mk_times (mk_number (Term.fastype_of t) k, t);
 
 (*Express t as a product of (possibly) a numeral with other sorted terms*)
-fun dest_coeff sign (Const (@{const_name Algebras.uminus}, _) $ t) = dest_coeff (~sign) t
+fun dest_coeff sign (Const (@{const_name Groups.uminus}, _) $ t) = dest_coeff (~sign) t
   | dest_coeff sign t =
     let val ts = sort TermOrd.term_ord (dest_prod t)
         val (n, ts') = find_first_numeral [] ts
@@ -104,7 +104,7 @@
   in mk_times (mk_divide (mk_number T p, mk_number T q), t) end;
 
 (*Express t as a product of a fraction with other sorted terms*)
-fun dest_fcoeff sign (Const (@{const_name Algebras.uminus}, _) $ t) = dest_fcoeff (~sign) t
+fun dest_fcoeff sign (Const (@{const_name Groups.uminus}, _) $ t) = dest_fcoeff (~sign) t
   | dest_fcoeff sign (Const (@{const_name Rings.divide}, _) $ t $ u) =
     let val (p, t') = dest_coeff sign t
         val (q, u') = dest_coeff 1 u
@@ -484,7 +484,7 @@
 in
 fun sign_conv pos_th neg_th ss t =
   let val T = fastype_of t;
-      val zero = Const(@{const_name Algebras.zero}, T);
+      val zero = Const(@{const_name Groups.zero}, T);
       val less = Const(@{const_name Orderings.less}, [T,T] ---> HOLogic.boolT);
       val pos = less $ zero $ t and neg = less $ t $ zero
       fun prove p =
--- a/src/HOL/Tools/record.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/record.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -145,16 +145,15 @@
       rep_inject RS named_cterm_instantiate [("abst", cterm_of typ_thy absC)] iso_tuple_intro;
     val (_, body) = Logic.dest_equals (List.last (prems_of intro_inst));
     val isomT = fastype_of body;
-    val isom_bind = Binding.name (name ^ isoN);
-    val isom_name = Sign.full_name typ_thy isom_bind;
+    val isom_binding = Binding.name (name ^ isoN);
+    val isom_name = Sign.full_name typ_thy isom_binding;
     val isom = Const (isom_name, isomT);
-    val isom_spec = (Thm.def_name (name ^ isoN), Logic.mk_equals (isom, body));
 
     val ([isom_def], cdef_thy) =
       typ_thy
-      |> Sign.add_consts_i [Syntax.no_syn (isom_bind, isomT)]
+      |> Sign.declare_const ((isom_binding, isomT), NoSyn) |> snd
       |> PureThy.add_defs false
-        [Thm.no_attributes (apfst (Binding.conceal o Binding.name) isom_spec)];
+        [((Binding.conceal (Thm.def_binding isom_binding), Logic.mk_equals (isom, body)), [])];
 
     val iso_tuple = isom_def RS (abs_inverse RS (rep_inject RS iso_tuple_intro));
     val cons = Const (@{const_name iso_tuple_cons}, isomT --> leftT --> rightT --> absT);
@@ -162,7 +161,7 @@
     val thm_thy =
       cdef_thy
       |> Iso_Tuple_Thms.map (Symtab.insert Thm.eq_thm_prop (isom_name, iso_tuple))
-      |> Sign.parent_path
+      |> Sign.restore_naming thy
       |> Code.add_default_eqn isom_def;
   in
     ((isom, cons $ isom), thm_thy)
@@ -280,11 +279,9 @@
 
 (* constructor *)
 
-fun mk_extC (name, T) Ts = (suffix extN name, Ts ---> T);
-
 fun mk_ext (name, T) ts =
   let val Ts = map fastype_of ts
-  in list_comb (Const (mk_extC (name, T) Ts), ts) end;
+  in list_comb (Const (suffix extN name, Ts ---> T), ts) end;
 
 
 (* selector *)
@@ -802,10 +799,7 @@
                   let
                     val (args, rest) = split_args (map fst (but_last fields)) fargs;
                     val more' = mk_ext rest;
-                  in
-                    (* FIXME authentic syntax *)
-                    list_comb (Syntax.const (suffix extN ext), args @ [more'])
-                  end
+                  in list_comb (Syntax.const (Syntax.mark_const (ext ^ extN)), args @ [more']) end
               | NONE => err ("no fields defined for " ^ ext))
           | NONE => err (name ^ " is no proper field"))
       | mk_ext [] = more;
@@ -848,96 +842,6 @@
 val print_record_type_abbr = Unsynchronized.ref true;
 val print_record_type_as_fields = Unsynchronized.ref true;
 
-local
-
-fun field_updates_tr' (tm as Const (c, _) $ k $ u) =
-      let
-        val t =
-          (case k of
-            Abs (_, _, Abs (_, _, t) $ Bound 0) =>
-              if null (loose_bnos t) then t else raise Match
-          | Abs (_, _, t) =>
-              if null (loose_bnos t) then t else raise Match
-          | _ => raise Match);
-      in
-        (case try (unsuffix updateN) c of
-          SOME name =>
-            (* FIXME check wrt. record data (??) *)
-            (* FIXME early extern (!??) *)
-            apfst (cons (Syntax.const @{syntax_const "_field_update"} $ Syntax.free name $ t))
-              (field_updates_tr' u)
-        | NONE => ([], tm))
-      end
-  | field_updates_tr' tm = ([], tm);
-
-fun record_update_tr' tm =
-  (case field_updates_tr' tm of
-    ([], _) => raise Match
-  | (ts, u) =>
-      Syntax.const @{syntax_const "_record_update"} $ u $
-        foldr1 (fn (v, w) => Syntax.const @{syntax_const "_field_updates"} $ v $ w) (rev ts));
-
-in
-
-fun field_update_tr' name =
-  let
-    (* FIXME authentic syntax *)
-    val update_name = suffix updateN name;
-    fun tr' [t, u] = record_update_tr' (Syntax.const update_name $ t $ u)
-      | tr' _ = raise Match;
-  in (update_name, tr') end;
-
-end;
-
-
-local
-
-(* FIXME early extern (!??) *)
-(* FIXME Syntax.free (??) *)
-fun field_tr' (c, t) = Syntax.const @{syntax_const "_field"} $ Syntax.const c $ t;
-
-fun fields_tr' (t, u) = Syntax.const @{syntax_const "_fields"} $ t $ u;
-
-fun record_tr' ctxt t =
-  let
-    val thy = ProofContext.theory_of ctxt;
-
-    fun strip_fields t =
-      (case strip_comb t of
-        (Const (ext, _), args as (_ :: _)) =>
-          (case try (unsuffix extN) (Sign.intern_const thy ext) of  (* FIXME authentic syntax *)
-            SOME ext' =>
-              (case get_extfields thy ext' of
-                SOME fields =>
-                 (let
-                    val f :: fs = but_last (map fst fields);
-                    val fields' = Sign.extern_const thy f :: map Long_Name.base_name fs;
-                    val (args', more) = split_last args;
-                  in (fields' ~~ args') @ strip_fields more end
-                  handle Library.UnequalLengths => [("", t)])
-              | NONE => [("", t)])
-          | NONE => [("", t)])
-       | _ => [("", t)]);
-
-    val (fields, (_, more)) = split_last (strip_fields t);
-    val _ = null fields andalso raise Match;
-    val u = foldr1 fields_tr' (map field_tr' fields);
-  in
-    case more of
-      Const (@{const_syntax Unity}, _) => Syntax.const @{syntax_const "_record"} $ u
-    | _ => Syntax.const @{syntax_const "_record_scheme"} $ u $ more
-  end;
-
-in
-
-fun record_ext_tr' name =
-  let
-    val ext_name = suffix extN name;
-    fun tr' ctxt ts = record_tr' ctxt (list_comb (Syntax.const ext_name, ts));
-  in (ext_name, tr') end;
-
-end;
-
 
 local
 
@@ -1059,6 +963,97 @@
 end;
 
 
+local
+
+(* FIXME Syntax.free (??) *)
+fun field_tr' (c, t) = Syntax.const @{syntax_const "_field"} $ Syntax.const c $ t;
+fun fields_tr' (t, u) = Syntax.const @{syntax_const "_fields"} $ t $ u;
+
+fun record_tr' ctxt t =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    val extern = Consts.extern (ProofContext.consts_of ctxt);
+
+    fun strip_fields t =
+      (case strip_comb t of
+        (Const (ext, _), args as (_ :: _)) =>
+          (case try (Syntax.unmark_const o unsuffix extN) ext of
+            SOME ext' =>
+              (case get_extfields thy ext' of
+                SOME fields =>
+                 (let
+                    val f :: fs = but_last (map fst fields);
+                    val fields' = extern f :: map Long_Name.base_name fs;
+                    val (args', more) = split_last args;
+                  in (fields' ~~ args') @ strip_fields more end
+                  handle Library.UnequalLengths => [("", t)])
+              | NONE => [("", t)])
+          | NONE => [("", t)])
+       | _ => [("", t)]);
+
+    val (fields, (_, more)) = split_last (strip_fields t);
+    val _ = null fields andalso raise Match;
+    val u = foldr1 fields_tr' (map field_tr' fields);
+  in
+    case more of
+      Const (@{const_syntax Unity}, _) => Syntax.const @{syntax_const "_record"} $ u
+    | _ => Syntax.const @{syntax_const "_record_scheme"} $ u $ more
+  end;
+
+in
+
+fun record_ext_tr' name =
+  let
+    val ext_name = Syntax.mark_const (name ^ extN);
+    fun tr' ctxt ts = record_tr' ctxt (list_comb (Syntax.const ext_name, ts));
+  in (ext_name, tr') end;
+
+end;
+
+
+local
+
+fun field_updates_tr' ctxt (tm as Const (c, _) $ k $ u) =
+      let
+        val extern = Consts.extern (ProofContext.consts_of ctxt);
+        val t =
+          (case k of
+            Abs (_, _, Abs (_, _, t) $ Bound 0) =>
+              if null (loose_bnos t) then t else raise Match
+          | Abs (_, _, t) =>
+              if null (loose_bnos t) then t else raise Match
+          | _ => raise Match);
+      in
+        (case Option.map extern (try Syntax.unmark_const c) of
+          SOME update_name =>
+            (case try (unsuffix updateN) update_name of
+              SOME name =>
+                apfst (cons (Syntax.const @{syntax_const "_field_update"} $ Syntax.free name $ t))
+                  (field_updates_tr' ctxt u)
+            | NONE => ([], tm))
+        | NONE => ([], tm))
+      end
+  | field_updates_tr' _ tm = ([], tm);
+
+fun record_update_tr' ctxt tm =
+  (case field_updates_tr' ctxt tm of
+    ([], _) => raise Match
+  | (ts, u) =>
+      Syntax.const @{syntax_const "_record_update"} $ u $
+        foldr1 (fn (v, w) => Syntax.const @{syntax_const "_field_updates"} $ v $ w) (rev ts));
+
+in
+
+fun field_update_tr' name =
+  let
+    val update_name = Syntax.mark_const (name ^ updateN);
+    fun tr' ctxt [t, u] = record_update_tr' ctxt (Syntax.const update_name $ t $ u)
+      | tr' _ _ = raise Match;
+  in (update_name, tr') end;
+
+end;
+
+
 
 (** record simprocs **)
 
@@ -1573,11 +1568,9 @@
           (s = @{const_name all} orelse s = @{const_name All}) andalso is_recT T
         | _ => false);
 
-    fun is_all t =
-      (case t of
-        Const (quantifier, _) $ _ =>
-          if quantifier = @{const_name all} orelse quantifier = @{const_name All} then ~1 else 0
-      | _ => 0);
+    fun is_all (Const (@{const_name all}, _) $ _) = ~1
+      | is_all (Const (@{const_name All}, _) $ _) = ~1
+      | is_all _ = 0;
   in
     if has_rec goal then
       Simplifier.full_simp_tac (HOL_basic_ss addsimprocs [record_split_simproc is_all]) i
@@ -1662,27 +1655,30 @@
 
 fun extension_definition name fields alphas zeta moreT more vars thy =
   let
-    val base = Long_Name.base_name;
+    val base_name = Long_Name.base_name name;
+
     val fieldTs = map snd fields;
+    val fields_moreTs = fieldTs @ [moreT];
+
     val alphas_zeta = alphas @ [zeta];
     val alphas_zetaTs = map (fn a => TFree (a, HOLogic.typeS)) alphas_zeta;
-    val extT_name = suffix ext_typeN name
-    val extT = Type (extT_name, alphas_zetaTs);
-    val fields_moreTs = fieldTs @ [moreT];
-
-
-    (*before doing anything else, create the tree of new types
-      that will back the record extension*)
+
+    val ext_binding = Binding.name (suffix extN base_name);
+    val ext_name = suffix extN name;
+    val extT = Type (suffix ext_typeN name, alphas_zetaTs);
+    val ext_type = fields_moreTs ---> extT;
+
+
+    (* the tree of new types that will back the record extension *)
 
     val mktreeV = Balanced_Tree.make Iso_Tuple_Support.mk_cons_tuple;
 
     fun mk_iso_tuple (left, right) (thy, i) =
       let
         val suff = if i = 0 then ext_typeN else inner_typeN ^ string_of_int i;
-        val nm = suffix suff (Long_Name.base_name name);
-        val ((_, cons), thy') =
-          Iso_Tuple_Support.add_iso_tuple_type
-            (nm, alphas_zeta) (fastype_of left, fastype_of right) thy;
+        val ((_, cons), thy') = thy
+          |> Iso_Tuple_Support.add_iso_tuple_type
+            (suffix suff base_name, alphas_zeta) (fastype_of left, fastype_of right);
       in
         (cons $ left $ right, (thy', i + 1))
       end;
@@ -1720,19 +1716,15 @@
 
     (* prepare declarations and definitions *)
 
-    (*fields constructor*)
-    val ext_decl = mk_extC (name, extT) fields_moreTs;
-    val ext_spec = list_comb (Const ext_decl, vars @ [more]) :== ext_body;
-
-    fun mk_ext args = list_comb (Const ext_decl, args);
-
-
     (* 1st stage part 2: define the ext constant *)
 
+    fun mk_ext args = list_comb (Const (ext_name, ext_type), args);
+    val ext_spec = Logic.mk_equals (mk_ext (vars @ [more]), ext_body);
+
     fun mk_defs () =
       typ_thy
-      |> Sign.add_consts_i [Syntax.no_syn (apfst (Binding.name o base) ext_decl)]
-      |> PureThy.add_defs false [Thm.no_attributes (apfst Binding.name ext_spec)]
+      |> Sign.declare_const ((ext_binding, ext_type), NoSyn) |> snd
+      |> PureThy.add_defs false [((Thm.def_binding ext_binding, ext_spec), [])]
       ||> Theory.checkpoint
     val ([ext_def], defs_thy) =
       timeit_msg "record extension constructor def:" mk_defs;
@@ -1856,7 +1848,7 @@
 fun obj_to_meta_all thm =
   let
     fun E thm =  (* FIXME proper name *)
-      (case (SOME (spec OF [thm]) handle THM _ => NONE) of
+      (case SOME (spec OF [thm]) handle THM _ => NONE of
         SOME thm' => E thm'
       | NONE => thm);
     val th1 = E thm;
@@ -1876,16 +1868,12 @@
 
 (* record_definition *)
 
-fun record_definition (args, b) parent (parents: parent_info list) raw_fields thy =
+fun record_definition (args, binding) parent (parents: parent_info list) raw_fields thy =
   let
-    val external_names = Name_Space.external_names (Sign.naming_of thy);
-
     val alphas = map fst args;
 
-    val base_name = Binding.name_of b;   (* FIXME include qualifier etc. (!?) *)
-    val name = Sign.full_name thy b;
-    val full = Sign.full_name_path thy base_name;
-    val base = Long_Name.base_name;
+    val name = Sign.full_name thy binding;
+    val full = Sign.full_name_path thy (Binding.name_of binding); (* FIXME Binding.qualified (!?) *)
 
     val bfields = map (fn (x, T, _) => (x, T)) raw_fields;
     val field_syntax = map #3 raw_fields;
@@ -1895,13 +1883,13 @@
     val parent_names = map fst parent_fields;
     val parent_types = map snd parent_fields;
     val parent_fields_len = length parent_fields;
-    val parent_variants = Name.variant_list [moreN, rN, rN ^ "'", wN] (map base parent_names);
+    val parent_variants =
+      Name.variant_list [moreN, rN, rN ^ "'", wN] (map Long_Name.base_name parent_names);
     val parent_vars = ListPair.map Free (parent_variants, parent_types);
     val parent_len = length parents;
 
     val fields = map (apfst full) bfields;
     val names = map fst fields;
-    val extN = full b;
     val types = map snd fields;
     val alphas_fields = fold Term.add_tfree_namesT types [];
     val alphas_ext = inter (op =) alphas_fields alphas;
@@ -1931,18 +1919,20 @@
     val all_named_vars_more = all_named_vars @ [(full_moreN, more)];
 
 
-    (* 1st stage: extension_thy *)
-
-    val ((extT, ext_induct, ext_inject, ext_surjective, ext_split, ext_def), extension_thy) =
+    (* 1st stage: ext_thy *)
+
+    val extension_name = full binding;
+
+    val ((extT, ext_induct, ext_inject, ext_surjective, ext_split, ext_def), ext_thy) =
       thy
-      |> Sign.add_path base_name
-      |> extension_definition extN fields alphas_ext zeta moreT more vars;
+      |> Sign.qualified_path false binding
+      |> extension_definition extension_name fields alphas_ext zeta moreT more vars;
 
     val _ = timing_msg "record preparing definitions";
     val Type extension_scheme = extT;
     val extension_name = unsuffix ext_typeN (fst extension_scheme);
     val extension = let val (n, Ts) = extension_scheme in (n, subst_last HOLogic.unitT Ts) end;
-    val extension_names = map (unsuffix ext_typeN o fst o #extension) parents @ [extN];
+    val extension_names = map (unsuffix ext_typeN o fst o #extension) parents @ [extension_name];
     val extension_id = implode extension_names;
 
     fun rec_schemeT n = mk_recordT (map #extension (drop n parents)) extT;
@@ -1978,13 +1968,9 @@
     val w = Free (wN, rec_schemeT 0)
 
 
-    (* prepare print translation functions *)
-    (* FIXME authentic syntax *)
-
-    val field_update_tr's =
-      map field_update_tr' (distinct (op =) (maps external_names (full_moreN :: names)));
-
-    val record_ext_tr's = map record_ext_tr' (external_names extN);
+    (* print translations *)
+
+    val external_names = Name_Space.external_names (Sign.naming_of ext_thy);
 
     val record_ext_type_abbr_tr's =
       let
@@ -1995,9 +1981,13 @@
     val record_ext_type_tr's =
       let
         (*avoid conflict with record_type_abbr_tr's*)
-        val trnames = if parent_len > 0 then external_names extN else [];
+        val trnames = if parent_len > 0 then external_names extension_name else [];
       in map record_ext_type_tr' trnames end;
 
+    val advanced_print_translation =
+      map field_update_tr' (full_moreN :: names) @ [record_ext_tr' extension_name] @
+      record_ext_type_tr's @ record_ext_type_abbr_tr's;
+
 
     (* prepare declarations *)
 
@@ -2013,8 +2003,8 @@
 
     (*record (scheme) type abbreviation*)
     val recordT_specs =
-      [(Binding.suffix_name schemeN b, alphas @ [zeta], rec_schemeT0, NoSyn),
-        (b, alphas, recT0, NoSyn)];
+      [(Binding.suffix_name schemeN binding, alphas @ [zeta], rec_schemeT0, NoSyn),
+        (binding, alphas, recT0, NoSyn)];
 
     val ext_defs = ext_def :: map #ext_def parents;
 
@@ -2035,8 +2025,8 @@
             fun to_Var (Free (c, T)) = Var ((c, 1), T);
           in mk_rec (map to_Var all_vars_more) 0 end;
 
-        val cterm_rec = cterm_of extension_thy r_rec0;
-        val cterm_vrs = cterm_of extension_thy r_rec0_Vars;
+        val cterm_rec = cterm_of ext_thy r_rec0;
+        val cterm_vrs = cterm_of ext_thy r_rec0_Vars;
         val insts = [("v", cterm_rec), ("v'", cterm_vrs)];
         val init_thm = named_cterm_instantiate insts updacc_eq_triv;
         val terminal = rtac updacc_eq_idI 1 THEN rtac refl 1;
@@ -2057,10 +2047,10 @@
     (*selectors*)
     fun mk_sel_spec ((c, T), thm) =
       let
-        val acc $ arg =
-          (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Envir.beta_eta_contract o concl_of) thm;
+        val (acc $ arg, _) =
+          HOLogic.dest_eq (HOLogic.dest_Trueprop (Envir.beta_eta_contract (Thm.concl_of thm)));
         val _ =
-          if (arg aconv r_rec0) then ()
+          if arg aconv r_rec0 then ()
           else raise TERM ("mk_sel_spec: different arg", [arg]);
       in
         Const (mk_selC rec_schemeT0 (c, T)) :== acc
@@ -2070,8 +2060,8 @@
     (*updates*)
     fun mk_upd_spec ((c, T), thm) =
       let
-        val upd $ _ $ arg =
-          fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (Envir.beta_eta_contract (concl_of thm))));
+        val (upd $ _ $ arg, _) =
+          HOLogic.dest_eq (HOLogic.dest_Trueprop (Envir.beta_eta_contract (Thm.concl_of thm)));
         val _ =
           if arg aconv r_rec0 then ()
           else raise TERM ("mk_sel_spec: different arg", [arg]);
@@ -2096,17 +2086,15 @@
     (* 2st stage: defs_thy *)
 
     fun mk_defs () =
-      extension_thy
-      |> Sign.add_trfuns ([], [], field_update_tr's, [])
-      |> Sign.add_advanced_trfuns
-        ([], [], record_ext_tr's @ record_ext_type_tr's @ record_ext_type_abbr_tr's, [])
-      |> Sign.parent_path
+      ext_thy
+      |> Sign.add_advanced_trfuns ([], [], advanced_print_translation, [])
+      |> Sign.restore_naming thy
       |> Sign.add_tyabbrs_i recordT_specs
-      |> Sign.add_path base_name
-      |> Sign.add_consts_i
-          (map2 (fn (x, T) => fn mx => (Binding.name x, T, mx)) sel_decls (field_syntax @ [NoSyn]))
-      |> (Sign.add_consts_i o map (fn (x, T) => (Binding.name x, T, NoSyn)))
-          (upd_decls @ [make_decl, fields_decl, extend_decl, truncate_decl])
+      |> Sign.qualified_path false binding
+      |> fold (fn ((x, T), mx) => snd o Sign.declare_const ((Binding.name x, T), mx))
+        (sel_decls ~~ (field_syntax @ [NoSyn]))
+      |> fold (fn (x, T) => snd o Sign.declare_const ((Binding.name x, T), NoSyn))
+        (upd_decls @ [make_decl, fields_decl, extend_decl, truncate_decl])
       |> (PureThy.add_defs false o map (Thm.no_attributes o apfst (Binding.conceal o Binding.name)))
         sel_specs
       ||>> (PureThy.add_defs false o map (Thm.no_attributes o apfst (Binding.conceal o Binding.name)))
@@ -2137,9 +2125,9 @@
     (*updates*)
     fun mk_upd_prop (i, (c, T)) =
       let
-        val x' = Free (Name.variant all_variants (base c ^ "'"), T --> T);
+        val x' = Free (Name.variant all_variants (Long_Name.base_name c ^ "'"), T --> T);
         val n = parent_fields_len + i;
-        val args' = nth_map n (K (x' $ nth all_vars_more n)) all_vars_more
+        val args' = nth_map n (K (x' $ nth all_vars_more n)) all_vars_more;
       in mk_upd updateN c x' r_rec0 === mk_rec args' 0 end;
     val upd_conv_props = ListPair.map mk_upd_prop (idxms, fields_more);
 
@@ -2401,7 +2389,7 @@
       |> add_record_splits extension_id (split_meta', split_object', split_ex', induct_scheme')
       |> add_extfields extension_name (fields @ [(full_moreN, moreT)])
       |> add_fieldext (extension_name, snd extension) (names @ [full_moreN])
-      |> Sign.parent_path;
+      |> Sign.restore_naming thy;
 
   in final_thy end;
 
@@ -2411,12 +2399,12 @@
 (*We do all preparations and error checks here, deferring the real
   work to record_definition.*)
 fun gen_add_record prep_typ prep_raw_parent quiet_mode
-    (params, b) raw_parent raw_fields thy =
+    (params, binding) raw_parent raw_fields thy =
   let
     val _ = Theory.requires thy "Record" "record definitions";
     val _ =
       if quiet_mode then ()
-      else writeln ("Defining record " ^ quote (Binding.name_of b) ^ " ...");
+      else writeln ("Defining record " ^ quote (Binding.str_of binding) ^ " ...");
 
     val ctxt = ProofContext.init thy;
 
@@ -2456,7 +2444,7 @@
 
     (* errors *)
 
-    val name = Sign.full_name thy b;
+    val name = Sign.full_name thy binding;
     val err_dup_record =
       if is_none (get_record thy name) then []
       else ["Duplicate definition of record " ^ quote name];
@@ -2493,9 +2481,9 @@
 
     val _ = if null errs then () else error (cat_lines errs);
   in
-    thy |> record_definition (args, b) parent parents bfields
+    thy |> record_definition (args, binding) parent parents bfields
   end
-  handle ERROR msg => cat_error msg ("Failed to define record " ^ quote (Binding.str_of b));
+  handle ERROR msg => cat_error msg ("Failed to define record " ^ quote (Binding.str_of binding));
 
 val add_record = gen_add_record cert_typ (K I);
 val add_record_cmd = gen_add_record read_typ read_raw_parent;
--- a/src/HOL/Tools/refute.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/refute.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -710,11 +710,11 @@
       | Const (@{const_name Finite_Set.finite}, _) => t
       | Const (@{const_name Orderings.less}, Type ("fun", [Type ("nat", []),
         Type ("fun", [Type ("nat", []), Type ("bool", [])])])) => t
-      | Const (@{const_name Algebras.plus}, Type ("fun", [Type ("nat", []),
+      | Const (@{const_name Groups.plus}, Type ("fun", [Type ("nat", []),
         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => t
-      | Const (@{const_name Algebras.minus}, Type ("fun", [Type ("nat", []),
+      | Const (@{const_name Groups.minus}, Type ("fun", [Type ("nat", []),
         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => t
-      | Const (@{const_name Algebras.times}, Type ("fun", [Type ("nat", []),
+      | Const (@{const_name Groups.times}, Type ("fun", [Type ("nat", []),
         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => t
       | Const (@{const_name List.append}, _) => t
       | Const (@{const_name lfp}, _) => t
@@ -886,13 +886,13 @@
       | Const (@{const_name Orderings.less}, T as Type ("fun", [Type ("nat", []),
         Type ("fun", [Type ("nat", []), Type ("bool", [])])])) =>
           collect_type_axioms T axs
-      | Const (@{const_name Algebras.plus}, T as Type ("fun", [Type ("nat", []),
+      | Const (@{const_name Groups.plus}, T as Type ("fun", [Type ("nat", []),
         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
           collect_type_axioms T axs
-      | Const (@{const_name Algebras.minus}, T as Type ("fun", [Type ("nat", []),
+      | Const (@{const_name Groups.minus}, T as Type ("fun", [Type ("nat", []),
         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
           collect_type_axioms T axs
-      | Const (@{const_name Algebras.times}, T as Type ("fun", [Type ("nat", []),
+      | Const (@{const_name Groups.times}, T as Type ("fun", [Type ("nat", []),
         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
           collect_type_axioms T axs
       | Const (@{const_name List.append}, T) => collect_type_axioms T axs
@@ -2794,7 +2794,7 @@
 
   fun Nat_plus_interpreter thy model args t =
     case t of
-      Const (@{const_name Algebras.plus}, Type ("fun", [Type ("nat", []),
+      Const (@{const_name Groups.plus}, Type ("fun", [Type ("nat", []),
         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
       let
         val size_of_nat = size_of_type thy model (Type ("nat", []))
@@ -2825,7 +2825,7 @@
 
   fun Nat_minus_interpreter thy model args t =
     case t of
-      Const (@{const_name Algebras.minus}, Type ("fun", [Type ("nat", []),
+      Const (@{const_name Groups.minus}, Type ("fun", [Type ("nat", []),
         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
       let
         val size_of_nat = size_of_type thy model (Type ("nat", []))
@@ -2853,7 +2853,7 @@
 
   fun Nat_times_interpreter thy model args t =
     case t of
-      Const (@{const_name Algebras.times}, Type ("fun", [Type ("nat", []),
+      Const (@{const_name Groups.times}, Type ("fun", [Type ("nat", []),
         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
       let
         val size_of_nat = size_of_type thy model (Type ("nat", []))
--- a/src/HOL/Tools/string_syntax.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/string_syntax.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -15,15 +15,14 @@
 
 (* nibble *)
 
-val nib_prefix = "String.nibble.";
-
 val mk_nib =
-  Syntax.Constant o unprefix nib_prefix o
+  Syntax.Constant o Syntax.mark_const o
     fst o Term.dest_Const o HOLogic.mk_nibble;
 
-fun dest_nib (Syntax.Constant c) =
-  HOLogic.dest_nibble (Const (nib_prefix ^ c, dummyT))
-    handle TERM _ => raise Match;
+fun dest_nib (Syntax.Constant s) =
+  (case try Syntax.unmark_const s of
+    NONE => raise Match
+  | SOME c => (HOLogic.dest_nibble (Const (c, HOLogic.nibbleT)) handle TERM _ => raise Match));
 
 
 (* char *)
--- a/src/HOL/Tools/typedef.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/Tools/typedef.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -119,8 +119,7 @@
       if def then
         theory
         |> Sign.add_consts_i [(name, setT', NoSyn)]   (* FIXME authentic syntax *)
-        |> PureThy.add_defs false
-          [((Binding.map_name Thm.def_name name, Logic.mk_equals (setC, set)), [])]
+        |> PureThy.add_defs false [((Thm.def_binding name, Logic.mk_equals (setC, set)), [])]
         |-> (fn [th] => pair (SOME th))
       else (NONE, theory);
     fun contract_def NONE th = th
--- a/src/HOL/UNITY/Comp/AllocBase.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/UNITY/Comp/AllocBase.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -58,9 +58,8 @@
 apply (rule monoI)
 apply (unfold prefix_def)
 apply (erule genPrefix.induct, auto)
-apply (simp add: union_le_mono)
 apply (erule order_trans)
-apply (rule union_upper1)
+apply simp
 done
 
 (** setsum **)
--- a/src/HOL/UNITY/Follows.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/UNITY/Follows.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -176,7 +176,7 @@
 apply (drule_tac x = "f xa" in spec)
 apply (drule_tac x = "g xa" in spec)
 apply (drule bspec, assumption) 
-apply (blast intro: union_le_mono order_trans)
+apply (blast intro: add_mono order_trans)
 done
 
 lemma Increasing_union: 
@@ -187,14 +187,14 @@
 apply (drule_tac x = "f xa" in spec)
 apply (drule_tac x = "g xa" in spec)
 apply (drule bspec, assumption) 
-apply (blast intro: union_le_mono order_trans)
+apply (blast intro: add_mono order_trans)
 done
 
 lemma Always_union:
      "[| F \<in> Always {s. f' s \<le> f s}; F \<in> Always {s. g' s \<le> g s} |]  
       ==> F \<in> Always {s. f' s + g' s \<le> f s + (g s :: ('a::order) multiset)}"
 apply (simp add: Always_eq_includes_reachable)
-apply (blast intro: union_le_mono)
+apply (blast intro: add_mono)
 done
 
 (*Except the last line, IDENTICAL to the proof script for Follows_Un_lemma*)
@@ -211,7 +211,7 @@
 apply (rule PSP_Stable)
 apply (erule_tac x = "f s" in spec)
 apply (erule Stable_Int, assumption, blast)
-apply (blast intro: union_le_mono order_trans)
+apply (blast intro: add_mono order_trans)
 done
 
 (*The !! is there to influence to effect of permutative rewriting at the end*)
--- a/src/HOL/ex/Arith_Examples.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/ex/Arith_Examples.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -33,7 +33,7 @@
 *)
 
 subsection {* Splitting of Operators: @{term max}, @{term min}, @{term abs},
-           @{term Algebras.minus}, @{term nat}, @{term Divides.mod},
+           @{term minus}, @{term nat}, @{term Divides.mod},
            @{term Divides.div} *}
 
 lemma "(i::nat) <= max i j"
--- a/src/HOL/ex/Binary.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/ex/Binary.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -24,8 +24,8 @@
     | dest_bit (Const (@{const_name True}, _)) = 1
     | dest_bit t = raise TERM ("dest_bit", [t]);
 
-  fun dest_binary (Const (@{const_name Algebras.zero}, Type (@{type_name nat}, _))) = 0
-    | dest_binary (Const (@{const_name Algebras.one}, Type (@{type_name nat}, _))) = 1
+  fun dest_binary (Const (@{const_name Groups.zero}, Type (@{type_name nat}, _))) = 0
+    | dest_binary (Const (@{const_name Groups.one}, Type (@{type_name nat}, _))) = 1
     | dest_binary (Const (@{const_name bit}, _) $ bs $ b) = 2 * dest_binary bs + dest_bit b
     | dest_binary t = raise TERM ("dest_binary", [t]);
 
@@ -191,7 +191,7 @@
 parse_translation {*
 let
   val syntax_consts =
-    map_aterms (fn Const (c, T) => Const (Syntax.constN ^ c, T) | a => a);
+    map_aterms (fn Const (c, T) => Const (Syntax.mark_const c, T) | a => a);
 
   fun binary_tr [Const (num, _)] =
         let
--- a/src/HOL/ex/SVC_Oracle.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/ex/SVC_Oracle.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -63,21 +63,21 @@
     (*abstraction of a numeric literal*)
     fun lit t = if can HOLogic.dest_number t then t else replace t;
     (*abstraction of a real/rational expression*)
-    fun rat ((c as Const(@{const_name Algebras.plus}, _)) $ x $ y) = c $ (rat x) $ (rat y)
-      | rat ((c as Const(@{const_name Algebras.minus}, _)) $ x $ y) = c $ (rat x) $ (rat y)
+    fun rat ((c as Const(@{const_name Groups.plus}, _)) $ x $ y) = c $ (rat x) $ (rat y)
+      | rat ((c as Const(@{const_name Groups.minus}, _)) $ x $ y) = c $ (rat x) $ (rat y)
       | rat ((c as Const(@{const_name Rings.divide}, _)) $ x $ y) = c $ (rat x) $ (rat y)
-      | rat ((c as Const(@{const_name Algebras.times}, _)) $ x $ y) = c $ (rat x) $ (rat y)
-      | rat ((c as Const(@{const_name Algebras.uminus}, _)) $ x) = c $ (rat x)
+      | rat ((c as Const(@{const_name Groups.times}, _)) $ x $ y) = c $ (rat x) $ (rat y)
+      | rat ((c as Const(@{const_name Groups.uminus}, _)) $ x) = c $ (rat x)
       | rat t = lit t
     (*abstraction of an integer expression: no div, mod*)
-    fun int ((c as Const(@{const_name Algebras.plus}, _)) $ x $ y) = c $ (int x) $ (int y)
-      | int ((c as Const(@{const_name Algebras.minus}, _)) $ x $ y) = c $ (int x) $ (int y)
-      | int ((c as Const(@{const_name Algebras.times}, _)) $ x $ y) = c $ (int x) $ (int y)
-      | int ((c as Const(@{const_name Algebras.uminus}, _)) $ x) = c $ (int x)
+    fun int ((c as Const(@{const_name Groups.plus}, _)) $ x $ y) = c $ (int x) $ (int y)
+      | int ((c as Const(@{const_name Groups.minus}, _)) $ x $ y) = c $ (int x) $ (int y)
+      | int ((c as Const(@{const_name Groups.times}, _)) $ x $ y) = c $ (int x) $ (int y)
+      | int ((c as Const(@{const_name Groups.uminus}, _)) $ x) = c $ (int x)
       | int t = lit t
     (*abstraction of a natural number expression: no minus*)
-    fun nat ((c as Const(@{const_name Algebras.plus}, _)) $ x $ y) = c $ (nat x) $ (nat y)
-      | nat ((c as Const(@{const_name Algebras.times}, _)) $ x $ y) = c $ (nat x) $ (nat y)
+    fun nat ((c as Const(@{const_name Groups.plus}, _)) $ x $ y) = c $ (nat x) $ (nat y)
+      | nat ((c as Const(@{const_name Groups.times}, _)) $ x $ y) = c $ (nat x) $ (nat y)
       | nat ((c as Const(@{const_name Suc}, _)) $ x) = c $ (nat x)
       | nat t = lit t
     (*abstraction of a relation: =, <, <=*)
--- a/src/HOL/ex/Summation.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/ex/Summation.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -28,7 +28,7 @@
 
 lemma \<Delta>_same_shift:
   assumes "\<Delta> f = \<Delta> g"
-  shows "\<exists>l. (op +) l \<circ> f = g"
+  shows "\<exists>l. plus l \<circ> f = g"
 proof -
   fix k
   from assms have "\<And>k. \<Delta> f k = \<Delta> g k" by simp
@@ -44,9 +44,9 @@
     show "f k - g k = f 0 - g 0"
       by (induct k rule: int_induct) (simp_all add: k_incr k_decr)
   qed
-  then have "\<And>k. ((op +) (g 0 - f 0) \<circ> f) k = g k"
+  then have "\<And>k. (plus (g 0 - f 0) \<circ> f) k = g k"
     by (simp add: algebra_simps)
-  then have "(op +) (g 0 - f 0) \<circ> f = g" ..
+  then have "plus (g 0 - f 0) \<circ> f = g" ..
   then show ?thesis ..
 qed
 
@@ -99,7 +99,7 @@
   "\<Sigma> (\<Delta> f) j l = f l - f j"
 proof -
   from \<Delta>_\<Sigma> have "\<Delta> (\<Sigma> (\<Delta> f) j) = \<Delta> f" .
-  then obtain k where "(op +) k \<circ> \<Sigma> (\<Delta> f) j = f" by (blast dest: \<Delta>_same_shift)
+  then obtain k where "plus k \<circ> \<Sigma> (\<Delta> f) j = f" by (blast dest: \<Delta>_same_shift)
   then have "\<And>q. f q = k + \<Sigma> (\<Delta> f) j q" by (simp add: expand_fun_eq)
   then show ?thesis by simp
 qed
--- a/src/HOL/ex/svc_funcs.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOL/ex/svc_funcs.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -146,16 +146,16 @@
     (*translation of a literal*)
     val lit = snd o HOLogic.dest_number;
     (*translation of a literal expression [no variables]*)
-    fun litExp (Const(@{const_name Algebras.plus}, T) $ x $ y) =
+    fun litExp (Const(@{const_name Groups.plus}, T) $ x $ y) =
           if is_numeric_op T then (litExp x) + (litExp y)
           else fail t
-      | litExp (Const(@{const_name Algebras.minus}, T) $ x $ y) =
+      | litExp (Const(@{const_name Groups.minus}, T) $ x $ y) =
           if is_numeric_op T then (litExp x) - (litExp y)
           else fail t
-      | litExp (Const(@{const_name Algebras.times}, T) $ x $ y) =
+      | litExp (Const(@{const_name Groups.times}, T) $ x $ y) =
           if is_numeric_op T then (litExp x) * (litExp y)
           else fail t
-      | litExp (Const(@{const_name Algebras.uminus}, T) $ x)   =
+      | litExp (Const(@{const_name Groups.uminus}, T) $ x)   =
           if is_numeric_op T then ~(litExp x)
           else fail t
       | litExp t = lit t
@@ -163,21 +163,21 @@
     (*translation of a real/rational expression*)
     fun suc t = Interp("+", [Int 1, t])
     fun tm (Const(@{const_name Suc}, T) $ x) = suc (tm x)
-      | tm (Const(@{const_name Algebras.plus}, T) $ x $ y) =
+      | tm (Const(@{const_name Groups.plus}, T) $ x $ y) =
           if is_numeric_op T then Interp("+", [tm x, tm y])
           else fail t
-      | tm (Const(@{const_name Algebras.minus}, T) $ x $ y) =
+      | tm (Const(@{const_name Groups.minus}, T) $ x $ y) =
           if is_numeric_op T then
               Interp("+", [tm x, Interp("*", [Int ~1, tm y])])
           else fail t
-      | tm (Const(@{const_name Algebras.times}, T) $ x $ y) =
+      | tm (Const(@{const_name Groups.times}, T) $ x $ y) =
           if is_numeric_op T then Interp("*", [tm x, tm y])
           else fail t
       | tm (Const(@{const_name Rings.inverse}, T) $ x) =
           if domain_type T = HOLogic.realT then
               Rat(1, litExp x)
           else fail t
-      | tm (Const(@{const_name Algebras.uminus}, T) $ x) =
+      | tm (Const(@{const_name Groups.uminus}, T) $ x) =
           if is_numeric_op T then Interp("*", [Int ~1, tm x])
           else fail t
       | tm t = Int (lit t)
--- a/src/HOLCF/IOA/meta_theory/Seq.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOLCF/IOA/meta_theory/Seq.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -8,7 +8,7 @@
 imports HOLCF
 begin
 
-domain 'a seq = nil | cons (HD :: 'a) (lazy TL :: "'a seq")  (infixr "##" 65)
+domain 'a seq = nil  ("nil") | cons (HD :: 'a) (lazy TL :: "'a seq")  (infixr "##" 65)
 
 (*
    sfilter       :: "('a -> tr) -> 'a seq -> 'a seq"
--- a/src/HOLCF/IOA/meta_theory/Sequence.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOLCF/IOA/meta_theory/Sequence.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -40,7 +40,7 @@
   "_partlist"     :: "args => 'a Seq"              ("[(_)?]")
 translations
   "[x, xs!]"     == "x>>[xs!]"
-  "[x!]"         == "x>>CONST nil"
+  "[x!]"         == "x>>nil"
   "[x, xs?]"     == "x>>[xs?]"
   "[x?]"         == "x>>CONST UU"
 
--- a/src/HOLCF/Tools/Domain/domain_syntax.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOLCF/Tools/Domain/domain_syntax.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -7,6 +7,7 @@
 signature DOMAIN_SYNTAX =
 sig
   val calc_syntax:
+      theory ->
       bool ->
       typ ->
       (string * typ list) *
@@ -28,7 +29,7 @@
 open Domain_Library;
 infixr 5 -->; infixr 6 ->>;
 
-fun calc_syntax  (* FIXME authentic syntax *)
+fun calc_syntax thy
     (definitional : bool)
     (dtypeprod : typ)
     ((dname : string, typevars : typ list), 
@@ -60,7 +61,7 @@
       val escape = let
         fun esc (c::cs) = if c mem ["'","_","(",")","/"] then "'"::c::esc cs
                           else      c::esc cs
-          |   esc []      = []
+          | esc []      = []
       in implode o esc o Symbol.explode end;
 
       fun dis_name_ con =
@@ -113,41 +114,45 @@
 
 (* ----- case translation --------------------------------------------------- *)
 
+    fun syntax b = Syntax.mark_const (Sign.full_bname thy b);
+
     local open Syntax in
     local
-      fun c_ast con mx = Constant (Binding.name_of con);   (* FIXME proper const syntax *)
-      fun expvar n     = Variable ("e"^(string_of_int n));
-      fun argvar n m _ = Variable ("a"^(string_of_int n)^"_"^
-                                   (string_of_int m));
+      fun c_ast authentic con = Constant ((authentic ? syntax) (Binding.name_of con));
+      fun expvar n = Variable ("e" ^ string_of_int n);
+      fun argvar n m _ = Variable ("a" ^ string_of_int n ^ "_" ^ string_of_int m);
       fun argvars n args = mapn (argvar n) 1 args;
-      fun app s (l,r)  = mk_appl (Constant s) [l,r];
+      fun app s (l, r) = mk_appl (Constant s) [l, r];
       val cabs = app "_cabs";
-      val capp = app "Rep_CFun";
-      fun con1 n (con,args,mx) = Library.foldl capp (c_ast con mx, argvars n args);
-      fun case1 n (con,args,mx) = app "_case1" (con1 n (con,args,mx), expvar n);
+      val capp = app @{const_syntax Rep_CFun};
+      fun con1 authentic n (con,args,mx) =
+        Library.foldl capp (c_ast authentic con, argvars n args);
+      fun case1 authentic n (con,args,mx) =
+        app "_case1" (con1 authentic n (con,args,mx), expvar n);
       fun arg1 n (con,args,_) = List.foldr cabs (expvar n) (argvars n args);
-      fun when1 n m = if n = m then arg1 n else K (Constant "UU");
+      fun when1 n m = if n = m then arg1 n else K (Constant @{const_syntax UU});
           
       fun app_var x = mk_appl (Constant "_variable") [x, Variable "rhs"];
       fun app_pat x = mk_appl (Constant "_pat") [x];
       fun args_list [] = Constant "_noargs"
-        |   args_list xs = foldr1 (app "_args") xs;
+        | args_list xs = foldr1 (app "_args") xs;
     in
-    val case_trans =
-        ParsePrintRule
-          (app "_case_syntax" (Variable "x", foldr1 (app "_case2") (mapn case1 1 cons')),
-           capp (Library.foldl capp (Constant (dnam^"_when"), mapn arg1 1 cons'), Variable "x"));
-        
-    fun one_abscon_trans n (con,mx,args) =
+    fun case_trans authentic =
         ParsePrintRule
-          (cabs (con1 n (con,mx,args), expvar n),
-           Library.foldl capp (Constant (dnam^"_when"), mapn (when1 n) 1 cons'));
-    val abscon_trans = mapn one_abscon_trans 1 cons';
+          (app "_case_syntax" (Variable "x", foldr1 (app "_case2") (mapn (case1 authentic) 1 cons')),
+           capp (Library.foldl capp
+            (Constant (syntax (dnam ^ "_when")), mapn arg1 1 cons'), Variable "x"));
         
-    fun one_case_trans (con,args,mx) =
+    fun one_abscon_trans authentic n (con,mx,args) =
+        ParsePrintRule
+          (cabs (con1 authentic n (con,mx,args), expvar n),
+           Library.foldl capp (Constant (syntax (dnam ^ "_when")), mapn (when1 n) 1 cons'));
+    fun abscon_trans authentic = mapn (one_abscon_trans authentic) 1 cons';
+        
+    fun one_case_trans authentic (con,args,mx) =
       let
-        val cname = c_ast con mx;
-        val pname = Constant (strip_esc (Binding.name_of con) ^ "_pat");
+        val cname = c_ast authentic con;
+        val pname = Constant (syntax (strip_esc (Binding.name_of con) ^ "_pat"));
         val ns = 1 upto length args;
         val xs = map (fn n => Variable ("x"^(string_of_int n))) ns;
         val ps = map (fn n => Variable ("p"^(string_of_int n))) ns;
@@ -160,7 +165,7 @@
          PrintRule (Library.foldl capp (cname, ListPair.map (app "_match") (ps,vs)),
                     app "_match" (mk_appl pname ps, args_list vs))]
         end;
-    val Case_trans = maps one_case_trans cons';
+    val Case_trans = maps (one_case_trans false) cons' @ maps (one_case_trans true) cons';
     end;
     end;
     val optional_consts =
@@ -169,7 +174,7 @@
   in (optional_consts @ [const_when] @ 
       consts_con @ consts_dis @ consts_mat @ consts_pat @ consts_sel @
       [const_take, const_finite],
-      (case_trans::(abscon_trans @ Case_trans)))
+      (case_trans false :: case_trans true :: (abscon_trans false @ abscon_trans true @ Case_trans)))
   end; (* let *)
 
 (* ----- putting all the syntax stuff together ------------------------------ *)
@@ -192,9 +197,9 @@
     val const_bisim =
         (Binding.name (comp_dnam^"_bisim"), relprod --> boolT, NoSyn);
     val ctt : ((binding * typ * mixfix) list * ast Syntax.trrule list) list =
-        map (calc_syntax definitional funprod) eqs';
+        map (calc_syntax thy'' definitional funprod) eqs';
   in thy''
-       |> ContConsts.add_consts_i
+       |> Cont_Consts.add_consts
            (maps fst ctt @ 
             (if length eqs'>1 andalso not definitional
              then [const_copy] else []) @
--- a/src/HOLCF/Tools/cont_consts.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/HOLCF/Tools/cont_consts.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -7,37 +7,28 @@
 
 signature CONT_CONSTS =
 sig
-  val add_consts: (binding * string * mixfix) list -> theory -> theory
-  val add_consts_i: (binding * typ * mixfix) list -> theory -> theory
+  val add_consts: (binding * typ * mixfix) list -> theory -> theory
+  val add_consts_cmd: (binding * string * mixfix) list -> theory -> theory
 end;
 
-structure ContConsts: CONT_CONSTS =
+structure Cont_Consts: CONT_CONSTS =
 struct
 
 
 (* misc utils *)
 
-fun first  (x,_,_) = x;
-fun second (_,x,_) = x;
-fun third  (_,_,x) = x;
-fun upd_first  f (x,y,z) = (f x,   y,   z);
-fun upd_second f (x,y,z) = (  x, f y,   z);
-fun upd_third  f (x,y,z) = (  x,   y, f z);
-
-fun change_arrow 0 T               = T
-|   change_arrow n (Type(_,[S,T])) = Type ("fun",[S,change_arrow (n-1) T])
-|   change_arrow _ _               = sys_error "cont_consts: change_arrow";
+fun change_arrow 0 T = T
+  | change_arrow n (Type (_, [S, T])) = Type ("fun", [S, change_arrow (n - 1) T])
+  | change_arrow _ T = raise TYPE ("cont_consts: change_arrow", [T], []);
 
 fun trans_rules name2 name1 n mx =
   let
-    fun argnames _ 0 = []
-    |   argnames c n = chr c::argnames (c+1) (n-1);
-    val vnames = argnames (ord "A") n;
+    val vnames = Name.invents Name.context "a" n;
     val extra_parse_rule = Syntax.ParseRule (Constant name2, Constant name1);
   in
     [Syntax.ParsePrintRule
       (Syntax.mk_appl (Constant name2) (map Variable vnames),
-        fold (fn arg => fn t => Syntax.mk_appl (Constant "Rep_CFun") [t, Variable arg])
+        fold (fn a => fn t => Syntax.mk_appl (Constant @{const_syntax Rep_CFun}) [t, Variable a])
           vnames (Constant name1))] @
     (case mx of
       Infix _ => [extra_parse_rule]
@@ -53,42 +44,51 @@
    declaration with the original name, type ...=>..., and the original mixfix
    is generated and connected to the other declaration via some translation.
 *)
-fun transform (c, T, mx) =
-    let
-        val c1 = Binding.name_of c;
-        val c2 = "_cont_" ^ c1;
-        val n  = Syntax.mixfix_args mx
-    in     ((c, T, NoSyn),
-            (Binding.name c2, change_arrow n T, mx),
-            trans_rules c2 c1 n mx) end;
+fun transform thy (c, T, mx) =
+  let
+    fun syntax b = Syntax.mark_const (Sign.full_bname thy b);
+    val c1 = Binding.name_of c;
+    val c2 = c1 ^ "_cont_syntax";
+    val n = Syntax.mixfix_args mx;
+  in
+    ((c, T, NoSyn),
+      (Binding.name c2, change_arrow n T, mx),
+      trans_rules (syntax c2) (syntax c1) n mx)
+  end;
 
-fun cfun_arity (Type(n,[_,T])) = if n = @{type_name "->"} then 1+cfun_arity T else 0
-|   cfun_arity _               = 0;
+fun cfun_arity (Type (n, [_, T])) = if n = @{type_name "->"} then 1 + cfun_arity T else 0
+  | cfun_arity _ = 0;
 
-fun is_contconst (_,_,NoSyn   ) = false
-|   is_contconst (_,_,Binder _) = false
-|   is_contconst (c,T,mx      ) =
-      cfun_arity T >= Syntax.mixfix_args mx
-        handle ERROR msg => cat_error msg ("in mixfix annotation for " ^ quote (Binding.str_of c));
+fun is_contconst (_, _, NoSyn) = false
+  | is_contconst (_, _, Binder _) = false    (* FIXME ? *)
+  | is_contconst (c, T, mx) =
+      let
+        val n = Syntax.mixfix_args mx handle ERROR msg =>
+          cat_error msg ("in mixfix annotation for " ^ quote (Binding.str_of c));
+      in cfun_arity T >= n end;
 
 
-(* add_consts(_i) *)
+(* add_consts *)
+
+local
 
 fun gen_add_consts prep_typ raw_decls thy =
   let
-    val decls = map (upd_second (prep_typ thy)) raw_decls;
+    val decls = map (fn (c, T, mx) => (c, prep_typ thy T, mx)) raw_decls;
     val (contconst_decls, normal_decls) = List.partition is_contconst decls;
-    val transformed_decls = map transform contconst_decls;
+    val transformed_decls = map (transform thy) contconst_decls;
   in
     thy
-    |> Sign.add_consts_i
-      (normal_decls @ map first transformed_decls @ map second transformed_decls)
-    (* FIXME authentic syntax *)
-    |> Sign.add_trrules_i (maps third transformed_decls)
+    |> Sign.add_consts_i (normal_decls @ map #1 transformed_decls @ map #2 transformed_decls)
+    |> Sign.add_trrules_i (maps #3 transformed_decls)
   end;
 
-val add_consts = gen_add_consts Syntax.read_typ_global;
-val add_consts_i = gen_add_consts Sign.certify_typ;
+in
+
+val add_consts = gen_add_consts Sign.certify_typ;
+val add_consts_cmd = gen_add_consts Syntax.read_typ_global;
+
+end;
 
 
 (* outer syntax *)
@@ -97,7 +97,7 @@
 
 val _ =
   OuterSyntax.command "consts" "declare constants (HOLCF)" K.thy_decl
-    (Scan.repeat1 P.const_binding >> (Toplevel.theory o add_consts));
+    (Scan.repeat1 P.const_binding >> (Toplevel.theory o add_consts_cmd));
 
 end;
 
--- a/src/Provers/Arith/abel_cancel.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Provers/Arith/abel_cancel.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -28,29 +28,29 @@
 
 (* FIXME dependent on abstract syntax *)
 
-fun zero t = Const (@{const_name Algebras.zero}, t);
-fun minus t = Const (@{const_name Algebras.uminus}, t --> t);
+fun zero t = Const (@{const_name Groups.zero}, t);
+fun minus t = Const (@{const_name Groups.uminus}, t --> t);
 
-fun add_terms pos (Const (@{const_name Algebras.plus}, _) $ x $ y, ts) =
+fun add_terms pos (Const (@{const_name Groups.plus}, _) $ x $ y, ts) =
       add_terms pos (x, add_terms pos (y, ts))
-  | add_terms pos (Const (@{const_name Algebras.minus}, _) $ x $ y, ts) =
+  | add_terms pos (Const (@{const_name Groups.minus}, _) $ x $ y, ts) =
       add_terms pos (x, add_terms (not pos) (y, ts))
-  | add_terms pos (Const (@{const_name Algebras.uminus}, _) $ x, ts) =
+  | add_terms pos (Const (@{const_name Groups.uminus}, _) $ x, ts) =
       add_terms (not pos) (x, ts)
   | add_terms pos (x, ts) = (pos,x) :: ts;
 
 fun terms fml = add_terms true (fml, []);
 
-fun zero1 pt (u as (c as Const(@{const_name Algebras.plus},_)) $ x $ y) =
+fun zero1 pt (u as (c as Const(@{const_name Groups.plus},_)) $ x $ y) =
       (case zero1 pt x of NONE => (case zero1 pt y of NONE => NONE
                                    | SOME z => SOME(c $ x $ z))
        | SOME z => SOME(c $ z $ y))
-  | zero1 (pos,t) (u as (c as Const(@{const_name Algebras.minus},_)) $ x $ y) =
+  | zero1 (pos,t) (u as (c as Const(@{const_name Groups.minus},_)) $ x $ y) =
       (case zero1 (pos,t) x of
          NONE => (case zero1 (not pos,t) y of NONE => NONE
                   | SOME z => SOME(c $ x $ z))
        | SOME z => SOME(c $ z $ y))
-  | zero1 (pos,t) (u as (c as Const(@{const_name Algebras.uminus},_)) $ x) =
+  | zero1 (pos,t) (u as (c as Const(@{const_name Groups.uminus},_)) $ x) =
       (case zero1 (not pos,t) x of NONE => NONE
        | SOME z => SOME(c $ z))
   | zero1 (pos,t) u =
@@ -71,7 +71,7 @@
 fun cancel t =
   let
     val c $ lhs $ rhs = t
-    val opp = case c of Const(@{const_name Algebras.plus},_) => true | _ => false;
+    val opp = case c of Const(@{const_name Groups.plus},_) => true | _ => false;
     val (pos,l) = find_common opp (terms lhs) (terms rhs)
     val posr = if opp then not pos else pos
     val t' = c $ (the(zero1 (pos,l) lhs)) $ (the(zero1 (posr,l) rhs))
--- a/src/Provers/Arith/cancel_div_mod.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Provers/Arith/cancel_div_mod.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -34,12 +34,12 @@
 functor CancelDivModFun(Data: CANCEL_DIV_MOD_DATA): CANCEL_DIV_MOD =
 struct
 
-fun coll_div_mod (Const(@{const_name Algebras.plus},_) $ s $ t) dms =
+fun coll_div_mod (Const(@{const_name Groups.plus},_) $ s $ t) dms =
       coll_div_mod t (coll_div_mod s dms)
-  | coll_div_mod (Const(@{const_name Algebras.times},_) $ m $ (Const(d,_) $ s $ n))
+  | coll_div_mod (Const(@{const_name Groups.times},_) $ m $ (Const(d,_) $ s $ n))
                  (dms as (divs,mods)) =
       if d = Data.div_name andalso m=n then ((s,n)::divs,mods) else dms
-  | coll_div_mod (Const(@{const_name Algebras.times},_) $ (Const(d,_) $ s $ n) $ m)
+  | coll_div_mod (Const(@{const_name Groups.times},_) $ (Const(d,_) $ s $ n) $ m)
                  (dms as (divs,mods)) =
       if d = Data.div_name andalso m=n then ((s,n)::divs,mods) else dms
   | coll_div_mod (Const(m,_) $ s $ n) (dms as (divs,mods)) =
@@ -55,8 +55,8 @@
    ==> thesis by transitivity
 *)
 
-val mk_plus = Data.mk_binop @{const_name Algebras.plus};
-val mk_times = Data.mk_binop @{const_name Algebras.times};
+val mk_plus = Data.mk_binop @{const_name Groups.plus};
+val mk_times = Data.mk_binop @{const_name Groups.times};
 
 fun rearrange t pq =
   let val ts = Data.dest_sum t;
--- a/src/Provers/trancl.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Provers/trancl.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -541,9 +541,11 @@
   val prems = flat (map_index (mkasm_trancl rel o swap) Hs);
   val prfs = solveTrancl (prems, C);
  in
-  Subgoal.FOCUS (fn {prems, ...} =>
-    let val thms = map (prove thy rel prems) prfs
-    in rtac (prove thy rel thms prf) 1 end) ctxt n st
+  Subgoal.FOCUS (fn {prems, concl, ...} =>
+    let
+      val SOME (_, _, rel', _) = decomp (term_of concl);
+      val thms = map (prove thy rel' prems) prfs
+    in rtac (prove thy rel' thms prf) 1 end) ctxt n st
  end
  handle Cannot => Seq.empty);
 
@@ -558,9 +560,11 @@
   val prems = flat (map_index (mkasm_rtrancl rel o swap) Hs);
   val prfs = solveRtrancl (prems, C);
  in
-  Subgoal.FOCUS (fn {prems, ...} =>
-    let val thms = map (prove thy rel prems) prfs
-    in rtac (prove thy rel thms prf) 1 end) ctxt n st
+  Subgoal.FOCUS (fn {prems, concl, ...} =>
+    let
+      val SOME (_, _, rel', _) = decomp (term_of concl);
+      val thms = map (prove thy rel' prems) prfs
+    in rtac (prove thy rel' thms prf) 1 end) ctxt n st
  end
  handle Cannot => Seq.empty | Subscript => Seq.empty);
 
--- a/src/Pure/Concurrent/task_queue.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/Concurrent/task_queue.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -34,12 +34,14 @@
   val finish: task -> queue -> bool * queue
 end;
 
-structure Task_Queue:> TASK_QUEUE =
+structure Task_Queue: TASK_QUEUE =
 struct
 
 (* tasks *)
 
-datatype task = Task of int option * serial;
+abstype task = Task of int option * serial
+with
+
 val dummy_task = Task (NONE, ~1);
 fun new_task pri = Task (pri, serial ());
 
@@ -47,15 +49,20 @@
 fun str_of_task (Task (_, i)) = string_of_int i;
 
 fun task_ord (Task t1, Task t2) = prod_ord (rev_order o option_ord int_ord) int_ord (t1, t2);
+val eq_task = is_equal o task_ord;
+
+end;
+
 structure Task_Graph = Graph(type key = task val ord = task_ord);
 
 
 (* nested groups *)
 
-datatype group = Group of
+abstype group = Group of
  {parent: group option,
   id: serial,
-  status: exn list Synchronized.var};
+  status: exn list Synchronized.var}
+with
 
 fun make_group (parent, id, status) = Group {parent = parent, id = id, status = status};
 
@@ -90,6 +97,8 @@
 fun str_of_group group =
   (is_canceled group ? enclose "(" ")") (string_of_int (group_id group));
 
+end;
+
 
 (* jobs *)
 
@@ -246,7 +255,7 @@
   let
     val group = get_group jobs task;
     val groups' = groups
-      |> fold (fn gid => Inttab.remove_list (op =) (gid, task)) (group_ancestry group);
+      |> fold (fn gid => Inttab.remove_list eq_task (gid, task)) (group_ancestry group);
     val jobs' = Task_Graph.del_node task jobs;
     val maximal = null (Task_Graph.imm_succs jobs task);
   in (maximal, make_queue groups' jobs') end;
--- a/src/Pure/Isar/expression.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/Isar/expression.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -604,7 +604,7 @@
 
 (* achieve plain syntax for locale predicates (without "PROP") *)
 
-fun aprop_tr' n c = (Syntax.constN ^ c, fn ctxt => fn args =>
+fun aprop_tr' n c = (Syntax.mark_const c, fn ctxt => fn args =>
   if length args = n then
     Syntax.const "_aprop" $   (* FIXME avoid old-style early externing (!??) *)
       Term.list_comb (Syntax.free (Consts.extern (ProofContext.consts_of ctxt) c), args)
@@ -641,8 +641,7 @@
       |> bodyT = propT ? Sign.add_advanced_trfuns ([], [], [aprop_tr' (length args) name], [])
       |> Sign.declare_const ((Binding.conceal binding, predT), NoSyn) |> snd
       |> PureThy.add_defs false
-        [((Binding.conceal (Binding.map_name Thm.def_name binding),
-            Logic.mk_equals (head, body)), [])];
+        [((Binding.conceal (Thm.def_binding binding), Logic.mk_equals (head, body)), [])];
     val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head;
 
     val cert = Thm.cterm_of defs_thy;
--- a/src/Pure/Isar/proof_context.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/Isar/proof_context.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -28,7 +28,6 @@
   val full_name: Proof.context -> binding -> string
   val syn_of: Proof.context -> Syntax.syntax
   val consts_of: Proof.context -> Consts.T
-  val const_syntax_name: Proof.context -> string -> string
   val the_const_constraint: Proof.context -> string -> typ
   val set_syntax_mode: Syntax.mode -> Proof.context -> Proof.context
   val restore_syntax_mode: Proof.context -> Proof.context -> Proof.context
@@ -235,7 +234,6 @@
 val restore_syntax_mode = map_syntax o Local_Syntax.restore_mode o syntax_of;
 
 val consts_of = #1 o #consts o rep_context;
-val const_syntax_name = Consts.syntax_name o consts_of;
 val the_const_constraint = Consts.the_constraint o consts_of;
 
 val facts_of = #facts o rep_context;
@@ -707,10 +705,10 @@
     val consts = consts_of ctxt;
   in
     t
-    |> Sign.extern_term (Consts.extern_early consts) thy
+    |> Sign.extern_term thy
     |> Local_Syntax.extern_term syntax
-    |> Syntax.standard_unparse_term (Consts.extern consts) ctxt (Local_Syntax.syn_of syntax)
-        (not (PureThy.old_appl_syntax thy))
+    |> Syntax.standard_unparse_term (Consts.extern consts) ctxt
+        (Local_Syntax.syn_of syntax) (not (PureThy.old_appl_syntax thy))
   end;
 
 in
@@ -990,7 +988,7 @@
 fun const_ast_tr intern ctxt [Syntax.Variable c] =
       let
         val Const (c', _) = read_const_proper ctxt c;
-        val d = if intern then const_syntax_name ctxt c' else c;
+        val d = if intern then Syntax.mark_const c' else c;
       in Syntax.Constant d end
   | const_ast_tr _ _ asts = raise Syntax.AST ("const_ast_tr", asts);
 
@@ -1018,7 +1016,9 @@
 
 fun const_syntax _ (Free (x, T), mx) = SOME (true, (x, T, mx))
   | const_syntax ctxt (Const (c, _), mx) =
-      Option.map (pair false) (try (Consts.syntax (consts_of ctxt)) (c, mx))
+      (case try (Consts.type_scheme (consts_of ctxt)) c of
+        SOME T => SOME (false, (Syntax.mark_const c, T, mx))
+      | NONE => NONE)
   | const_syntax _ _ = NONE;
 
 in
--- a/src/Pure/ML/ml_antiquote.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/ML/ml_antiquote.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -114,7 +114,7 @@
 
 fun const syn = Args.context -- Scan.lift Args.name_source >> (fn (ctxt, c) =>
   #1 (Term.dest_Const (ProofContext.read_const_proper ctxt c))
-  |> syn ? ProofContext.const_syntax_name ctxt
+  |> syn ? Syntax.mark_const
   |> ML_Syntax.print_string);
 
 val _ = inline "const_name" (const false);
--- a/src/Pure/Proof/proof_syntax.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/Proof/proof_syntax.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -67,8 +67,8 @@
        ("", paramT --> paramsT, Delimfix "_")]
   |> Sign.add_modesyntax_i (Symbol.xsymbolsN, true)
       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1\\<Lambda>_./ _)", [0, 3], 3)),
-       (Syntax.constN ^ "Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ \\<cdot>/ _)", [4, 5], 4)),
-       (Syntax.constN ^ "AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ \\<bullet>/ _)", [4, 5], 4))]
+       (Syntax.mark_const "Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ \\<cdot>/ _)", [4, 5], 4)),
+       (Syntax.mark_const "AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ \\<bullet>/ _)", [4, 5], 4))]
   |> Sign.add_modesyntax_i ("latex", false)
       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1\\<^bold>\\<lambda>_./ _)", [0, 3], 3))]
   |> Sign.add_trrules_i (map Syntax.ParsePrintRule
@@ -78,10 +78,10 @@
           [Variable "l", Syntax.mk_appl (Constant "_Lam") [Variable "m", Variable "A"]]),
        (Syntax.mk_appl (Constant "_Lam")
           [Syntax.mk_appl (Constant "_Lam1") [Variable "x", Variable "A"], Variable "B"],
-        Syntax.mk_appl (Constant (Syntax.constN ^ "AbsP")) [Variable "A",
+        Syntax.mk_appl (Constant (Syntax.mark_const "AbsP")) [Variable "A",
           (Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "B"])]),
        (Syntax.mk_appl (Constant "_Lam") [Variable "x", Variable "A"],
-        Syntax.mk_appl (Constant (Syntax.constN ^ "Abst"))
+        Syntax.mk_appl (Constant (Syntax.mark_const "Abst"))
           [(Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "A"])])]);
 
 
--- a/src/Pure/Syntax/lexicon.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/Syntax/lexicon.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -31,7 +31,11 @@
   val read_xnum: string -> {radix: int, leading_zeros: int, value: int}
   val read_float: string -> {mant: int, exp: int}
   val fixedN: string
+  val mark_fixed: string -> string
+  val unmark_fixed: string -> string
   val constN: string
+  val mark_const: string -> string
+  val unmark_const: string -> string
 end;
 
 signature LEXICON =
@@ -331,8 +335,13 @@
 
 (* specific identifiers *)
 
+val fixedN = "\\<^fixed>";
+val mark_fixed = prefix fixedN;
+val unmark_fixed = unprefix fixedN;
+
 val constN = "\\<^const>";
-val fixedN = "\\<^fixed>";
+val mark_const = prefix constN;
+val unmark_const = unprefix constN;
 
 
 (* read numbers *)
--- a/src/Pure/Syntax/printer.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/Syntax/printer.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -321,10 +321,10 @@
             else pr, args))
 
     and atomT a =
-      (case try (unprefix Lexicon.constN) a of
+      (case try Lexicon.unmark_const a of
         SOME c => Pretty.mark (Markup.const c) (Pretty.str (extern_const c))
       | NONE =>
-          (case try (unprefix Lexicon.fixedN) a of
+          (case try Lexicon.unmark_fixed a of
             SOME x => the (token_trans "_free" x)
           | NONE => Pretty.str a))
 
@@ -340,8 +340,8 @@
       let
         val nargs = length args;
         val markup = Pretty.mark
-          (Markup.const (unprefix Lexicon.constN a) handle Fail _ =>
-            (Markup.fixed (unprefix Lexicon.fixedN a)))
+          (Markup.const (Lexicon.unmark_const a) handle Fail _ =>
+            (Markup.fixed (Lexicon.unmark_fixed a)))
           handle Fail _ => I;
 
         (*find matching table entry, or print as prefix / postfix*)
--- a/src/Pure/Syntax/syn_trans.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/Syntax/syn_trans.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -129,13 +129,15 @@
 
 (* type propositions *)
 
-fun mk_type ty = Lexicon.const "_constrain" $ Lexicon.const "TYPE" $ (Lexicon.const "itself" $ ty);
+fun mk_type ty =
+  Lexicon.const "_constrain" $
+    Lexicon.const "\\<^const>TYPE" $ (Lexicon.const "itself" $ ty);
 
 fun ofclass_tr (*"_ofclass"*) [ty, cls] = cls $ mk_type ty
   | ofclass_tr (*"_ofclass"*) ts = raise TERM ("ofclass_tr", ts);
 
 fun sort_constraint_tr (*"_sort_constraint"*) [ty] =
-      Lexicon.const "Pure.sort_constraint" $ mk_type ty
+      Lexicon.const "\\<^const>Pure.sort_constraint" $ mk_type ty
   | sort_constraint_tr (*"_sort_constraint"*) ts = raise TERM ("sort_constraint_tr", ts);
 
 
@@ -152,7 +154,7 @@
         (case Ast.unfold_ast_p "_asms" asms of
           (asms', Ast.Appl [Ast.Constant "_asm", asm']) => asms' @ [asm']
         | _ => raise Ast.AST ("bigimpl_ast_tr", asts))
-      in Ast.fold_ast_p "==>" (prems, concl) end
+      in Ast.fold_ast_p "\\<^const>==>" (prems, concl) end
   | bigimpl_ast_tr (*"_bigimpl"*) asts = raise Ast.AST ("bigimpl_ast_tr", asts);
 
 
@@ -350,7 +352,7 @@
 
 (* type propositions *)
 
-fun type_prop_tr' _ (*"_type_prop"*) T [Const ("Pure.sort_constraint", _)] =
+fun type_prop_tr' _ (*"_type_prop"*) T [Const ("\\<^const>Pure.sort_constraint", _)] =
       Lexicon.const "_sort_constraint" $ TypeExt.term_of_typ true T
   | type_prop_tr' show_sorts (*"_type_prop"*) T [t] =
       Lexicon.const "_ofclass" $ TypeExt.term_of_typ show_sorts T $ t
@@ -366,7 +368,7 @@
     fun is_prop Ts t =
       fastype_of1 (Ts, t) = propT handle TERM _ => false;
 
-    fun is_term (Const ("Pure.term", _) $ _) = true
+    fun is_term (Const ("\\<^const>Pure.term", _) $ _) = true
       | is_term _ = false;
 
     fun tr' _ (t as Const _) = t
@@ -379,7 +381,7 @@
       | tr' Ts (t as Bound _) =
           if is_prop Ts t then aprop t else t
       | tr' Ts (Abs (x, T, t)) = Abs (x, T, tr' (T :: Ts) t)
-      | tr' Ts (t as t1 $ (t2 as Const ("TYPE", Type ("itself", [T])))) =
+      | tr' Ts (t as t1 $ (t2 as Const ("\\<^const>TYPE", Type ("itself", [T])))) =
           if is_prop Ts t andalso not (is_term t) then Const ("_type_prop", T) $ tr' Ts t1
           else tr' Ts t1 $ tr' Ts t2
       | tr' Ts (t as t1 $ t2) =
@@ -393,7 +395,7 @@
 fun impl_ast_tr' (*"==>"*) asts =
   if TypeExt.no_brackets () then raise Match
   else
-    (case Ast.unfold_ast_p "==>" (Ast.Appl (Ast.Constant "==>" :: asts)) of
+    (case Ast.unfold_ast_p "\\<^const>==>" (Ast.Appl (Ast.Constant "\\<^const>==>" :: asts)) of
       (prems as _ :: _ :: _, concl) =>
         let
           val (asms, asm) = split_last prems;
@@ -514,11 +516,13 @@
    [("_abs", abs_ast_tr'),
     ("_idts", idtyp_ast_tr' "_idts"),
     ("_pttrns", idtyp_ast_tr' "_pttrns"),
-    ("==>", impl_ast_tr'),
+    ("\\<^const>==>", impl_ast_tr'),
     ("_index", index_ast_tr')]);
 
 val pure_trfunsT =
-  [("_type_prop", type_prop_tr'), ("TYPE", type_tr'), ("_type_constraint_", type_constraint_tr')];
+ [("_type_prop", type_prop_tr'),
+  ("\\<^const>TYPE", type_tr'),
+  ("_type_constraint_", type_constraint_tr')];
 
 fun struct_trfuns structs =
   ([], [("_struct", struct_tr structs)], [], [("_struct", struct_ast_tr' structs)]);
--- a/src/Pure/Syntax/syntax.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/Syntax/syntax.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -302,7 +302,7 @@
       lexicon =
         if changed then fold Scan.extend_lexicon (SynExt.delims_of xprods) lexicon else lexicon,
       gram = if changed then Parser.extend_gram gram xprods else gram,
-      consts = Library.merge (op =) (consts1, filter_out Long_Name.is_qualified consts2),
+      consts = Library.merge (op =) (consts1, filter_out (can Lexicon.unmark_const) consts2),
       prmodes = insert (op =) mode (Library.merge (op =) (prmodes1, prmodes2)),
       parse_ast_trtab =
         update_trtab "parse ast translation" (if_inout parse_ast_translation) parse_ast_trtab,
--- a/src/Pure/Syntax/type_ext.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/Syntax/type_ext.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -123,7 +123,9 @@
       | decode (t $ u) = decode t $ decode u
       | decode (Const (a, T)) =
           let val c =
-            (case try (unprefix Lexicon.constN) a of SOME c => c | NONE => snd (map_const a))
+            (case try Lexicon.unmark_const a of
+              SOME c => c
+            | NONE => snd (map_const a))
           in Const (c, map_type T) end
       | decode (Free (a, T)) =
           (case (map_free a, map_const a) of
--- a/src/Pure/axclass.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/axclass.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -474,7 +474,7 @@
     val ([def], def_thy) =
       thy
       |> Sign.primitive_class (bclass, super)
-      |> PureThy.add_defs false [((Binding.map_name Thm.def_name bconst, class_eq), [])];
+      |> PureThy.add_defs false [((Thm.def_binding bconst, class_eq), [])];
     val (raw_intro, (raw_classrel, raw_axioms)) =
       split_defined (length conjs) def ||> chop (length super);
 
--- a/src/Pure/consts.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/consts.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -21,15 +21,13 @@
   val space_of: T -> Name_Space.T
   val is_concealed: T -> string -> bool
   val intern: T -> xstring -> string
+  val intern_syntax: T -> xstring -> string
   val extern: T -> string -> xstring
-  val extern_early: T -> string -> xstring
-  val syntax: T -> string * mixfix -> string * typ * mixfix
-  val syntax_name: T -> string -> string
   val read_const: T -> string -> term
   val certify: Pretty.pp -> Type.tsig -> bool -> T -> term -> term  (*exception TYPE*)
   val typargs: T -> string * typ -> typ list
   val instance: T -> string * typ list -> typ
-  val declare: bool -> Name_Space.naming -> binding * typ -> T -> T
+  val declare: Name_Space.naming -> binding * typ -> T -> T
   val constrain: string * typ option -> T -> T
   val abbreviate: Pretty.pp -> Type.tsig -> Name_Space.naming -> string ->
     binding * term -> T -> (term * term) * T
@@ -46,7 +44,7 @@
 
 (* datatype T *)
 
-type decl = {T: typ, typargs: int list list, authentic: bool};
+type decl = {T: typ, typargs: int list list};
 type abbrev = {rhs: term, normal_rhs: term, force_expand: bool};
 
 datatype T = Consts of
@@ -128,18 +126,10 @@
 val intern = Name_Space.intern o space_of;
 val extern = Name_Space.extern o space_of;
 
-fun extern_early consts c =
-  (case try (the_const consts) c of
-    SOME ({authentic = true, ...}, _) => Syntax.constN ^ c
-  | _ => extern consts c);
-
-fun syntax consts (c, mx) =
-  let
-    val ({T, authentic, ...}, _) = the_const consts c handle TYPE (msg, _, _) => error msg;
-    val c' = if authentic then Syntax.constN ^ c else Long_Name.base_name c;
-  in (c', T, mx) end;
-
-fun syntax_name consts c = #1 (syntax consts (c, NoSyn));
+fun intern_syntax consts name =
+  (case try Syntax.unmark_const name of
+    SOME c => c
+  | NONE => intern consts name);
 
 
 (* read_const *)
@@ -231,10 +221,10 @@
 
 (* declarations *)
 
-fun declare authentic naming (b, declT) =
+fun declare naming (b, declT) =
   map_consts (fn (decls, constraints, rev_abbrevs) =>
     let
-      val decl = {T = declT, typargs = typargs_of declT, authentic = authentic};
+      val decl = {T = declT, typargs = typargs_of declT};
       val (_, decls') = decls |> Name_Space.define true naming (b, (decl, NONE));
     in (decls', constraints, rev_abbrevs) end);
 
@@ -285,7 +275,7 @@
   in
     consts |> map_consts (fn (decls, constraints, rev_abbrevs) =>
       let
-        val decl = {T = T, typargs = typargs_of T, authentic = true};
+        val decl = {T = T, typargs = typargs_of T};
         val abbr = {rhs = rhs, normal_rhs = normal_rhs, force_expand = force_expand};
         val (_, decls') = decls
           |> Name_Space.define true naming (b, (decl, SOME abbr));
--- a/src/Pure/drule.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/drule.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -77,7 +77,6 @@
   val flexflex_unique: thm -> thm
   val export_without_context: thm -> thm
   val export_without_context_open: thm -> thm
-  val get_def: theory -> xstring -> thm
   val store_thm: binding -> thm -> thm
   val store_standard_thm: binding -> thm -> thm
   val store_thm_open: binding -> thm -> thm
@@ -451,8 +450,6 @@
 
 val read_prop = certify o Simple_Syntax.read_prop;
 
-fun get_def thy = Thm.axiom thy o Name_Space.intern (Theory.axiom_space thy) o Thm.def_name;
-
 fun store_thm name th =
   Context.>>> (Context.map_theory_result (PureThy.store_thm (name, th)));
 
--- a/src/Pure/more_thm.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/more_thm.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -77,6 +77,7 @@
   val untag: string -> attribute
   val def_name: string -> string
   val def_name_optional: string -> string -> string
+  val def_binding: Binding.binding -> Binding.binding
   val def_binding_optional: Binding.binding -> Binding.binding -> Binding.binding
   val has_name_hint: thm -> bool
   val get_name_hint: thm -> string
@@ -384,8 +385,10 @@
 fun def_name_optional c "" = def_name c
   | def_name_optional _ name = name;
 
+val def_binding = Binding.map_name def_name;
+
 fun def_binding_optional b name =
-  if Binding.is_empty name then Binding.map_name def_name b else name;
+  if Binding.is_empty name then def_binding b else name;
 
 
 (* unofficial theorem names *)
--- a/src/Pure/old_goals.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/old_goals.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -16,6 +16,7 @@
   val simple_read_term: theory -> typ -> string -> term
   val read_term: theory -> string -> term
   val read_prop: theory -> string -> term
+  val get_def: theory -> xstring -> thm
   type proof
   val premises: unit -> thm list
   val reset_goals: unit -> unit
@@ -220,6 +221,8 @@
 fun read_prop thy = simple_read_term thy propT;
 
 
+fun get_def thy = Thm.axiom thy o Name_Space.intern (Theory.axiom_space thy) o Thm.def_name;
+
 
 (*** Goal package ***)
 
--- a/src/Pure/pure_thy.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/pure_thy.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -225,6 +225,7 @@
 
 val typ = Simple_Syntax.read_typ;
 val prop = Simple_Syntax.read_prop;
+val const = Syntax.mark_const;
 
 val typeT = Syntax.typeT;
 val spropT = Syntax.spropT;
@@ -310,11 +311,11 @@
     ("_indexvar",   typ "index",                       Delimfix "'\\<index>"),
     ("_struct",     typ "index => logic",              Mixfix ("\\<struct>_", [1000], 1000)),
     ("_update_name", typ "idt",                        NoSyn),
-    ("==>",         typ "prop => prop => prop",        Delimfix "op ==>"),
-    (Term.dummy_patternN, typ "aprop",                 Delimfix "'_"),
+    (const "==>",   typ "prop => prop => prop",        Delimfix "op ==>"),
+    (const Term.dummy_patternN, typ "aprop",           Delimfix "'_"),
     ("_sort_constraint", typ "type => prop",           Delimfix "(1SORT'_CONSTRAINT/(1'(_')))"),
-    ("Pure.term",   typ "logic => prop",               Delimfix "TERM _"),
-    ("Pure.conjunction", typ "prop => prop => prop",   Infixr ("&&&", 2))]
+    (const "Pure.term", typ "logic => prop",           Delimfix "TERM _"),
+    (const "Pure.conjunction", typ "prop => prop => prop", Infixr ("&&&", 2))]
   #> Sign.add_syntax_i applC_syntax
   #> Sign.add_modesyntax_i (Symbol.xsymbolsN, true)
    [("fun",      typ "type => type => type",   Mixfix ("(_/ \\<Rightarrow> _)", [1, 0], 0)),
@@ -326,14 +327,14 @@
     ("_idtypdummy", typ "type => idt",         Mixfix ("'_()\\<Colon>_", [], 0)),
     ("_type_constraint_", typ "'a",            NoSyn),
     ("_lambda",  typ "pttrns => 'a => logic",  Mixfix ("(3\\<lambda>_./ _)", [0, 3], 3)),
-    ("==",       typ "'a => 'a => prop",       Infixr ("\\<equiv>", 2)),
-    ("all_binder", typ "idts => prop => prop", Mixfix ("(3\\<And>_./ _)", [0, 0], 0)),
-    ("==>",      typ "prop => prop => prop",   Infixr ("\\<Longrightarrow>", 1)),
+    (const "==", typ "'a => 'a => prop",       Infixr ("\\<equiv>", 2)),
+    (const "all_binder", typ "idts => prop => prop", Mixfix ("(3\\<And>_./ _)", [0, 0], 0)),
+    (const "==>", typ "prop => prop => prop",  Infixr ("\\<Longrightarrow>", 1)),
     ("_DDDOT",   typ "aprop",                  Delimfix "\\<dots>"),
     ("_bigimpl", typ "asms => prop => prop",   Mixfix ("((1\\<lbrakk>_\\<rbrakk>)/ \\<Longrightarrow> _)", [0, 1], 1)),
     ("_DDDOT",   typ "logic",                  Delimfix "\\<dots>")]
   #> Sign.add_modesyntax_i ("", false)
-   [("prop", typ "prop => prop", Mixfix ("_", [0], 0))]
+   [(const "prop", typ "prop => prop", Mixfix ("_", [0], 0))]
   #> Sign.add_modesyntax_i ("HTML", false)
    [("_lambda", typ "pttrns => 'a => logic", Mixfix ("(3\\<lambda>_./ _)", [0, 3], 3))]
   #> Sign.add_consts_i
--- a/src/Pure/sign.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Pure/sign.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -41,7 +41,6 @@
   val declared_tyname: theory -> string -> bool
   val declared_const: theory -> string -> bool
   val const_monomorphic: theory -> string -> bool
-  val const_syntax_name: theory -> string -> string
   val const_typargs: theory -> string * typ -> typ list
   val const_instance: theory -> string * typ list -> typ
   val mk_const: theory -> string * typ list -> term
@@ -59,7 +58,7 @@
   val intern_typ: theory -> typ -> typ
   val extern_typ: theory -> typ -> typ
   val intern_term: theory -> term -> term
-  val extern_term: (string -> xstring) -> theory -> term -> term
+  val extern_term: theory -> term -> term
   val intern_tycons: theory -> typ -> typ
   val arity_number: theory -> string -> int
   val arity_sorts: theory -> string -> sort -> sort list
@@ -226,7 +225,6 @@
 val the_const_type = Consts.the_type o consts_of;
 val const_type = try o the_const_type;
 val const_monomorphic = Consts.is_monomorphic o consts_of;
-val const_syntax_name = Consts.syntax_name o consts_of;
 val const_typargs = Consts.typargs o consts_of;
 val const_instance = Consts.instance o consts_of;
 
@@ -299,7 +297,7 @@
 val intern_typ = typ_mapping intern_class intern_type;
 val extern_typ = typ_mapping extern_class extern_type;
 val intern_term = term_mapping intern_class intern_type intern_const;
-fun extern_term h = term_mapping extern_class extern_type (K h);
+val extern_term = term_mapping extern_class extern_type (K Syntax.mark_const);
 val intern_tycons = typ_mapping (K I) intern_type;
 
 end;
@@ -486,7 +484,10 @@
 fun notation add mode args thy =
   let
     val change_gram = if add then Syntax.update_const_gram else Syntax.remove_const_gram;
-    fun const_syntax (Const (c, _), mx) = try (Consts.syntax (consts_of thy)) (c, mx)
+    fun const_syntax (Const (c, _), mx) =
+          (case try (Consts.type_scheme (consts_of thy)) c of
+            SOME T => SOME (Syntax.mark_const c, T, mx)
+          | NONE => NONE)
       | const_syntax _ = NONE;
   in gen_syntax change_gram (K I) mode (map_filter const_syntax args) thy end;
 
@@ -495,35 +496,34 @@
 
 local
 
-fun gen_add_consts parse_typ authentic raw_args thy =
+fun gen_add_consts parse_typ raw_args thy =
   let
     val ctxt = ProofContext.init thy;
     val prepT = Type.no_tvars o Term.no_dummyT o certify_typ thy o parse_typ ctxt;
     fun prep (b, raw_T, mx) =
       let
         val c = full_name thy b;
-        val c_syn = if authentic then Syntax.constN ^ c else Name.of_binding b;
         val T = (prepT raw_T handle TYPE (msg, _, _) => error msg) handle ERROR msg =>
           cat_error msg ("in declaration of constant " ^ quote (Binding.str_of b));
         val T' = Logic.varifyT T;
-      in ((b, T'), (c_syn, T', mx), Const (c, T)) end;
+      in ((b, T'), (Syntax.mark_const c, T', mx), Const (c, T)) end;
     val args = map prep raw_args;
   in
     thy
-    |> map_consts (fold (Consts.declare authentic (naming_of thy) o #1) args)
+    |> map_consts (fold (Consts.declare (naming_of thy) o #1) args)
     |> add_syntax_i (map #2 args)
     |> pair (map #3 args)
   end;
 
 in
 
-fun add_consts args = snd o gen_add_consts Syntax.parse_typ false args;
-fun add_consts_i args = snd o gen_add_consts (K I) false args;
+fun add_consts args = snd o gen_add_consts Syntax.parse_typ args;
+fun add_consts_i args = snd o gen_add_consts (K I) args;
 
 fun declare_const ((b, T), mx) thy =
   let
     val pos = Binding.pos_of b;
-    val ([const as Const (c, _)], thy') = gen_add_consts (K I) true [(b, T, mx)] thy;
+    val ([const as Const (c, _)], thy') = gen_add_consts (K I) [(b, T, mx)] thy;
     val _ = Position.report (Markup.const_decl c) pos;
   in (const, thy') end;
 
--- a/src/Sequents/ILL_predlog.thy	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/Sequents/ILL_predlog.thy	Mon Feb 22 11:19:15 2010 -0800
@@ -11,7 +11,7 @@
   disj :: "[plf,plf] => plf"   (infixr "|" 35)
   impl :: "[plf,plf] => plf"   (infixr ">" 35)
   eq :: "[plf,plf] => plf"   (infixr "=" 35)
-  ff    :: "plf"
+  ff :: "plf"    ("ff")
 
   PL    :: "plf => o"      ("[* / _ / *]" [] 100)
 
@@ -22,8 +22,8 @@
 
   "[* A & B *]" == "[*A*] && [*B*]"
   "[* A | B *]" == "![*A*] ++ ![*B*]"
-  "[* - A *]"   == "[*A > CONST ff*]"
-  "[* XCONST ff *]" == "0"
+  "[* - A *]" == "[*A > ff*]"
+  "[* ff *]" == "0"
   "[* A = B *]" => "[* (A > B) & (B > A) *]"
 
   "[* A > B *]" == "![*A*] -o [*B*]"
--- a/src/ZF/Tools/datatype_package.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/ZF/Tools/datatype_package.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -298,7 +298,7 @@
 
   (*** Prove the recursor theorems ***)
 
-  val recursor_eqns = case try (Drule.get_def thy1) recursor_base_name of
+  val recursor_eqns = case try (OldGoals.get_def thy1) recursor_base_name of
      NONE => (writeln "  [ No recursion operator ]";
               [])
    | SOME recursor_def =>
--- a/src/ZF/Tools/inductive_package.ML	Mon Feb 22 11:17:41 2010 -0800
+++ b/src/ZF/Tools/inductive_package.ML	Mon Feb 22 11:19:15 2010 -0800
@@ -179,7 +179,7 @@
 
   (*fetch fp definitions from the theory*)
   val big_rec_def::part_rec_defs =
-    map (Drule.get_def thy1)
+    map (OldGoals.get_def thy1)
         (case rec_names of [_] => rec_names
                          | _   => big_rec_base_name::rec_names);