merge
authorhellerar
Wed, 01 Sep 2010 17:19:47 +0200
changeset 39093 84af4fdc1a98
parent 39080 cae59dc0a094 (current diff)
parent 38974 e109feb514a8 (diff)
child 39094 67da17aced5a
merge
src/Pure/ProofGeneral/proof_general_keywords.ML
--- a/Admin/Benchmarks/HOL-datatype/ROOT.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/Admin/Benchmarks/HOL-datatype/ROOT.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -5,11 +5,11 @@
 
 val tests = ["Brackin", "Instructions", "SML", "Verilog"];
 
-Unsynchronized.set timing;
+timing := true;
 
-warning "\nset quick_and_dirty\n"; Unsynchronized.set quick_and_dirty;
+warning "\nset quick_and_dirty\n"; quick_and_dirty := true;
 use_thys tests;
 
-warning "\nreset quick_and_dirty\n"; Unsynchronized.reset quick_and_dirty;
+warning "\nreset quick_and_dirty\n"; quick_and_dirty := false;
 List.app Thy_Info.remove_thy tests;
 use_thys tests;
--- a/Admin/Benchmarks/HOL-record/ROOT.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/Admin/Benchmarks/HOL-record/ROOT.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -5,10 +5,10 @@
 
 val tests = ["RecordBenchmark"];
 
-Unsynchronized.set timing;
+timing := true;
 
-warning "\nset quick_and_dirty\n"; Unsynchronized.set quick_and_dirty;
+warning "\nset quick_and_dirty\n"; quick_and_dirty := true;
 use_thys tests;
 
-warning "\nreset quick_and_dirty\n"; Unsynchronized.reset quick_and_dirty;
+warning "\nreset quick_and_dirty\n"; quick_and_dirty := false;
 use_thys tests;
--- a/Admin/Benchmarks/HOL-record/RecordBenchmark.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/Admin/Benchmarks/HOL-record/RecordBenchmark.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -8,7 +8,7 @@
 imports Main
 begin
 
-ML {* Unsynchronized.set Record.timing *}
+ML {* Record.timing := true *}
 
 record many_A =
 A000::nat
--- a/Admin/update-keywords	Thu Aug 26 13:15:37 2010 +0200
+++ b/Admin/update-keywords	Wed Sep 01 17:19:47 2010 +0200
@@ -11,9 +11,9 @@
 cd "$ISABELLE_HOME/etc"
 
 isabelle keywords \
-  "$LOG/Pure.gz" "$LOG/Pure-ProofGeneral.gz" "$LOG/HOL.gz" "$LOG/HOLCF.gz" \
-  "$LOG/HOL-Boogie.gz" "$LOG/HOL-Nominal.gz" "$LOG/HOL-Statespace.gz"
+  "$LOG/Pure.gz" "$LOG/HOL.gz" "$LOG/HOLCF.gz" "$LOG/HOL-Boogie.gz" \
+  "$LOG/HOL-Nominal.gz" "$LOG/HOL-Statespace.gz"
 
 isabelle keywords -k ZF \
-  "$LOG/Pure.gz" "$LOG/Pure-ProofGeneral.gz" "$LOG/FOL.gz" "$LOG/ZF.gz"
+  "$LOG/Pure.gz" "$LOG/FOL.gz" "$LOG/ZF.gz"
 
--- a/NEWS	Thu Aug 26 13:15:37 2010 +0200
+++ b/NEWS	Wed Sep 01 17:19:47 2010 +0200
@@ -23,6 +23,28 @@
 at the cost of clarity of file dependencies.  Recall that Isabelle/ML
 files exclusively use the .ML extension.  Minor INCOMPATIBILTY.
 
+* Various options that affect document antiquotations are now properly
+handled within the context via configuration options, instead of
+unsynchronized references.  There are both ML Config.T entities and
+Isar declaration attributes to access these.
+
+  ML:                       Isar:
+
+  Thy_Output.display        thy_output_display
+  Thy_Output.quotes         thy_output_quotes
+  Thy_Output.indent         thy_output_indent
+  Thy_Output.source         thy_output_source
+  Thy_Output.break          thy_output_break
+
+Note that the corresponding "..._default" references may be only
+changed globally at the ROOT session setup, but *not* within a theory.
+
+* ML structure Unsynchronized never opened, not even in Isar
+interaction mode as before.  Old Unsynchronized.set etc. have been
+discontinued -- use plain := instead.  This should be *rare* anyway,
+since modern tools always work via official context data, notably
+configuration options.
+
 
 *** Pure ***
 
@@ -40,6 +62,13 @@
 
 *** HOL ***
 
+* Renamed class eq and constant eq (for code generation) to class equal
+and constant equal, plus renaming of related facts and various tuning.
+INCOMPATIBILITY.
+
+* Scala (2.8 or higher) has been added to the target languages of
+the code generator.
+
 * Dropped type classes mult_mono and mult_mono1.  INCOMPATIBILITY.
 
 * Theory SetsAndFunctions has been split into Function_Algebras and Set_Algebras;
@@ -104,6 +133,10 @@
     Trueprop ~> HOL.Trueprop
     True ~> HOL.True
     False ~> HOL.False
+    op & ~> HOL.conj
+    op | ~> HOL.disj
+    op --> ~> HOL.implies
+    op = ~> HOL.eq
     Not ~> HOL.Not
     The ~> HOL.The
     All ~> HOL.All
--- a/doc-src/Classes/Thy/Classes.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Classes/Thy/Classes.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -8,14 +8,14 @@
   Type classes were introduced by Wadler and Blott \cite{wadler89how}
   into the Haskell language to allow for a reasonable implementation
   of overloading\footnote{throughout this tutorial, we are referring
-  to classical Haskell 1.0 type classes, not considering
-  later additions in expressiveness}.
-  As a canonical example, a polymorphic equality function
-  @{text "eq \<Colon> \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"} which is overloaded on different
-  types for @{text "\<alpha>"}, which is achieved by splitting introduction
-  of the @{text eq} function from its overloaded definitions by means
-  of @{text class} and @{text instance} declarations:
-  \footnote{syntax here is a kind of isabellized Haskell}
+  to classical Haskell 1.0 type classes, not considering later
+  additions in expressiveness}.  As a canonical example, a polymorphic
+  equality function @{text "eq \<Colon> \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"} which is overloaded on
+  different types for @{text "\<alpha>"}, which is achieved by splitting
+  introduction of the @{text eq} function from its overloaded
+  definitions by means of @{text class} and @{text instance}
+  declarations: \footnote{syntax here is a kind of isabellized
+  Haskell}
 
   \begin{quote}
 
@@ -41,14 +41,14 @@
   these annotations are assertions that a particular polymorphic type
   provides definitions for overloaded functions.
 
-  Indeed, type classes not only allow for simple overloading
-  but form a generic calculus, an instance of order-sorted
-  algebra \cite{nipkow-sorts93,Nipkow-Prehofer:1993,Wenzel:1997:TPHOL}.
+  Indeed, type classes not only allow for simple overloading but form
+  a generic calculus, an instance of order-sorted algebra
+  \cite{nipkow-sorts93,Nipkow-Prehofer:1993,Wenzel:1997:TPHOL}.
 
-  From a software engineering point of view, type classes
-  roughly correspond to interfaces in object-oriented languages like Java;
-  so, it is naturally desirable that type classes do not only
-  provide functions (class parameters) but also state specifications
+  From a software engineering point of view, type classes roughly
+  correspond to interfaces in object-oriented languages like Java; so,
+  it is naturally desirable that type classes do not only provide
+  functions (class parameters) but also state specifications
   implementations must obey.  For example, the @{text "class eq"}
   above could be given the following specification, demanding that
   @{text "class eq"} is an equivalence relation obeying reflexivity,
@@ -65,11 +65,10 @@
 
   \end{quote}
 
-  \noindent From a theoretical point of view, type classes are lightweight
-  modules; Haskell type classes may be emulated by
-  SML functors \cite{classes_modules}. 
-  Isabelle/Isar offers a discipline of type classes which brings
-  all those aspects together:
+  \noindent From a theoretical point of view, type classes are
+  lightweight modules; Haskell type classes may be emulated by SML
+  functors \cite{classes_modules}.  Isabelle/Isar offers a discipline
+  of type classes which brings all those aspects together:
 
   \begin{enumerate}
     \item specifying abstract parameters together with
@@ -81,15 +80,15 @@
       locales \cite{kammueller-locales}.
   \end{enumerate}
 
-  \noindent Isar type classes also directly support code generation
-  in a Haskell like fashion. Internally, they are mapped to more primitive 
-  Isabelle concepts \cite{Haftmann-Wenzel:2006:classes}.
+  \noindent Isar type classes also directly support code generation in
+  a Haskell like fashion. Internally, they are mapped to more
+  primitive Isabelle concepts \cite{Haftmann-Wenzel:2006:classes}.
 
-  This tutorial demonstrates common elements of structured specifications
-  and abstract reasoning with type classes by the algebraic hierarchy of
-  semigroups, monoids and groups.  Our background theory is that of
-  Isabelle/HOL \cite{isa-tutorial}, for which some
-  familiarity is assumed.
+  This tutorial demonstrates common elements of structured
+  specifications and abstract reasoning with type classes by the
+  algebraic hierarchy of semigroups, monoids and groups.  Our
+  background theory is that of Isabelle/HOL \cite{isa-tutorial}, for
+  which some familiarity is assumed.
 *}
 
 section {* A simple algebra example \label{sec:example} *}
@@ -107,25 +106,24 @@
   assumes assoc: "(x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"
 
 text {*
-  \noindent This @{command class} specification consists of two
-  parts: the \qn{operational} part names the class parameter
-  (@{element "fixes"}), the \qn{logical} part specifies properties on them
-  (@{element "assumes"}).  The local @{element "fixes"} and
-  @{element "assumes"} are lifted to the theory toplevel,
-  yielding the global
+  \noindent This @{command class} specification consists of two parts:
+  the \qn{operational} part names the class parameter (@{element
+  "fixes"}), the \qn{logical} part specifies properties on them
+  (@{element "assumes"}).  The local @{element "fixes"} and @{element
+  "assumes"} are lifted to the theory toplevel, yielding the global
   parameter @{term [source] "mult \<Colon> \<alpha>\<Colon>semigroup \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>"} and the
-  global theorem @{fact "semigroup.assoc:"}~@{prop [source] "\<And>x y
-  z \<Colon> \<alpha>\<Colon>semigroup. (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"}.
+  global theorem @{fact "semigroup.assoc:"}~@{prop [source] "\<And>x y z \<Colon>
+  \<alpha>\<Colon>semigroup. (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"}.
 *}
 
 
 subsection {* Class instantiation \label{sec:class_inst} *}
 
 text {*
-  The concrete type @{typ int} is made a @{class semigroup}
-  instance by providing a suitable definition for the class parameter
-  @{text "(\<otimes>)"} and a proof for the specification of @{fact assoc}.
-  This is accomplished by the @{command instantiation} target:
+  The concrete type @{typ int} is made a @{class semigroup} instance
+  by providing a suitable definition for the class parameter @{text
+  "(\<otimes>)"} and a proof for the specification of @{fact assoc}.  This is
+  accomplished by the @{command instantiation} target:
 *}
 
 instantiation %quote int :: semigroup
@@ -143,22 +141,22 @@
 end %quote
 
 text {*
-  \noindent @{command instantiation} defines class parameters
-  at a particular instance using common specification tools (here,
-  @{command definition}).  The concluding @{command instance}
-  opens a proof that the given parameters actually conform
-  to the class specification.  Note that the first proof step
-  is the @{method default} method,
-  which for such instance proofs maps to the @{method intro_classes} method.
-  This reduces an instance judgement to the relevant primitive
-  proof goals; typically it is the first method applied
-  in an instantiation proof.
+  \noindent @{command instantiation} defines class parameters at a
+  particular instance using common specification tools (here,
+  @{command definition}).  The concluding @{command instance} opens a
+  proof that the given parameters actually conform to the class
+  specification.  Note that the first proof step is the @{method
+  default} method, which for such instance proofs maps to the @{method
+  intro_classes} method.  This reduces an instance judgement to the
+  relevant primitive proof goals; typically it is the first method
+  applied in an instantiation proof.
 
-  From now on, the type-checker will consider @{typ int}
-  as a @{class semigroup} automatically, i.e.\ any general results
-  are immediately available on concrete instances.
+  From now on, the type-checker will consider @{typ int} as a @{class
+  semigroup} automatically, i.e.\ any general results are immediately
+  available on concrete instances.
 
-  \medskip Another instance of @{class semigroup} yields the natural numbers:
+  \medskip Another instance of @{class semigroup} yields the natural
+  numbers:
 *}
 
 instantiation %quote nat :: semigroup
@@ -177,21 +175,20 @@
 end %quote
 
 text {*
-  \noindent Note the occurence of the name @{text mult_nat}
-  in the primrec declaration;  by default, the local name of
-  a class operation @{text f} to be instantiated on type constructor
-  @{text \<kappa>} is mangled as @{text f_\<kappa>}.  In case of uncertainty,
-  these names may be inspected using the @{command "print_context"} command
-  or the corresponding ProofGeneral button.
+  \noindent Note the occurence of the name @{text mult_nat} in the
+  primrec declaration; by default, the local name of a class operation
+  @{text f} to be instantiated on type constructor @{text \<kappa>} is
+  mangled as @{text f_\<kappa>}.  In case of uncertainty, these names may be
+  inspected using the @{command "print_context"} command or the
+  corresponding ProofGeneral button.
 *}
 
 subsection {* Lifting and parametric types *}
 
 text {*
-  Overloaded definitions given at a class instantiation
-  may include recursion over the syntactic structure of types.
-  As a canonical example, we model product semigroups
-  using our simple algebra:
+  Overloaded definitions given at a class instantiation may include
+  recursion over the syntactic structure of types.  As a canonical
+  example, we model product semigroups using our simple algebra:
 *}
 
 instantiation %quote prod :: (semigroup, semigroup) semigroup
@@ -211,21 +208,19 @@
 text {*
   \noindent Associativity of product semigroups is established using
   the definition of @{text "(\<otimes>)"} on products and the hypothetical
-  associativity of the type components;  these hypotheses
-  are legitimate due to the @{class semigroup} constraints imposed
-  on the type components by the @{command instance} proposition.
-  Indeed, this pattern often occurs with parametric types
-  and type classes.
+  associativity of the type components; these hypotheses are
+  legitimate due to the @{class semigroup} constraints imposed on the
+  type components by the @{command instance} proposition.  Indeed,
+  this pattern often occurs with parametric types and type classes.
 *}
 
 
 subsection {* Subclassing *}
 
 text {*
-  We define a subclass @{text monoidl} (a semigroup with a left-hand neutral)
-  by extending @{class semigroup}
-  with one additional parameter @{text neutral} together
-  with its characteristic property:
+  We define a subclass @{text monoidl} (a semigroup with a left-hand
+  neutral) by extending @{class semigroup} with one additional
+  parameter @{text neutral} together with its characteristic property:
 *}
 
 class %quote monoidl = semigroup +
@@ -233,10 +228,10 @@
   assumes neutl: "\<one> \<otimes> x = x"
 
 text {*
-  \noindent Again, we prove some instances, by
-  providing suitable parameter definitions and proofs for the
-  additional specifications.  Observe that instantiations
-  for types with the same arity may be simultaneous:
+  \noindent Again, we prove some instances, by providing suitable
+  parameter definitions and proofs for the additional specifications.
+  Observe that instantiations for types with the same arity may be
+  simultaneous:
 *}
 
 instantiation %quote nat and int :: monoidl
@@ -309,8 +304,8 @@
 end %quote
 
 text {*
-  \noindent To finish our small algebra example, we add a @{text group} class
-  with a corresponding instance:
+  \noindent To finish our small algebra example, we add a @{text
+  group} class with a corresponding instance:
 *}
 
 class %quote group = monoidl +
@@ -338,9 +333,9 @@
 subsection {* A look behind the scenes *}
 
 text {*
-  The example above gives an impression how Isar type classes work
-  in practice.  As stated in the introduction, classes also provide
-  a link to Isar's locale system.  Indeed, the logical core of a class
+  The example above gives an impression how Isar type classes work in
+  practice.  As stated in the introduction, classes also provide a
+  link to Isar's locale system.  Indeed, the logical core of a class
   is nothing other than a locale:
 *}
 
@@ -402,13 +397,14 @@
 qed
 
 text {*
-  \noindent Here the \qt{@{keyword "in"} @{class group}} target specification
-  indicates that the result is recorded within that context for later
-  use.  This local theorem is also lifted to the global one @{fact
-  "group.left_cancel:"} @{prop [source] "\<And>x y z \<Colon> \<alpha>\<Colon>group. x \<otimes> y = x \<otimes>
-  z \<longleftrightarrow> y = z"}.  Since type @{text "int"} has been made an instance of
-  @{text "group"} before, we may refer to that fact as well: @{prop
-  [source] "\<And>x y z \<Colon> int. x \<otimes> y = x \<otimes> z \<longleftrightarrow> y = z"}.
+  \noindent Here the \qt{@{keyword "in"} @{class group}} target
+  specification indicates that the result is recorded within that
+  context for later use.  This local theorem is also lifted to the
+  global one @{fact "group.left_cancel:"} @{prop [source] "\<And>x y z \<Colon>
+  \<alpha>\<Colon>group. x \<otimes> y = x \<otimes> z \<longleftrightarrow> y = z"}.  Since type @{text "int"} has been
+  made an instance of @{text "group"} before, we may refer to that
+  fact as well: @{prop [source] "\<And>x y z \<Colon> int. x \<otimes> y = x \<otimes> z \<longleftrightarrow> y =
+  z"}.
 *}
 
 
@@ -424,15 +420,14 @@
 
 text {*
   \noindent If the locale @{text group} is also a class, this local
-  definition is propagated onto a global definition of
-  @{term [source] "pow_nat \<Colon> nat \<Rightarrow> \<alpha>\<Colon>monoid \<Rightarrow> \<alpha>\<Colon>monoid"}
-  with corresponding theorems
+  definition is propagated onto a global definition of @{term [source]
+  "pow_nat \<Colon> nat \<Rightarrow> \<alpha>\<Colon>monoid \<Rightarrow> \<alpha>\<Colon>monoid"} with corresponding theorems
 
   @{thm pow_nat.simps [no_vars]}.
 
-  \noindent As you can see from this example, for local
-  definitions you may use any specification tool
-  which works together with locales, such as Krauss's recursive function package
+  \noindent As you can see from this example, for local definitions
+  you may use any specification tool which works together with
+  locales, such as Krauss's recursive function package
   \cite{krauss2006}.
 *}
 
@@ -440,19 +435,17 @@
 subsection {* A functor analogy *}
 
 text {*
-  We introduced Isar classes by analogy to type classes in
-  functional programming;  if we reconsider this in the
-  context of what has been said about type classes and locales,
-  we can drive this analogy further by stating that type
-  classes essentially correspond to functors that have
-  a canonical interpretation as type classes.
-  There is also the possibility of other interpretations.
-  For example, @{text list}s also form a monoid with
-  @{text append} and @{term "[]"} as operations, but it
-  seems inappropriate to apply to lists
-  the same operations as for genuinely algebraic types.
-  In such a case, we can simply make a particular interpretation
-  of monoids for lists:
+  We introduced Isar classes by analogy to type classes in functional
+  programming; if we reconsider this in the context of what has been
+  said about type classes and locales, we can drive this analogy
+  further by stating that type classes essentially correspond to
+  functors that have a canonical interpretation as type classes.
+  There is also the possibility of other interpretations.  For
+  example, @{text list}s also form a monoid with @{text append} and
+  @{term "[]"} as operations, but it seems inappropriate to apply to
+  lists the same operations as for genuinely algebraic types.  In such
+  a case, we can simply make a particular interpretation of monoids
+  for lists:
 *}
 
 interpretation %quote list_monoid: monoid append "[]"
@@ -510,12 +503,10 @@
 qed
 
 text {*
-  The logical proof is carried out on the locale level.
-  Afterwards it is propagated
-  to the type system, making @{text group} an instance of
-  @{text monoid} by adding an additional edge
-  to the graph of subclass relations
-  (\figref{fig:subclass}).
+  The logical proof is carried out on the locale level.  Afterwards it
+  is propagated to the type system, making @{text group} an instance
+  of @{text monoid} by adding an additional edge to the graph of
+  subclass relations (\figref{fig:subclass}).
 
   \begin{figure}[htbp]
    \begin{center}
@@ -547,8 +538,8 @@
    \end{center}
   \end{figure}
 
-  For illustration, a derived definition
-  in @{text group} using @{text pow_nat}
+  For illustration, a derived definition in @{text group} using @{text
+  pow_nat}
 *}
 
 definition %quote (in group) pow_int :: "int \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>" where
@@ -557,17 +548,17 @@
     else (pow_nat (nat (- k)) x)\<div>)"
 
 text {*
-  \noindent yields the global definition of
-  @{term [source] "pow_int \<Colon> int \<Rightarrow> \<alpha>\<Colon>group \<Rightarrow> \<alpha>\<Colon>group"}
-  with the corresponding theorem @{thm pow_int_def [no_vars]}.
+  \noindent yields the global definition of @{term [source] "pow_int \<Colon>
+  int \<Rightarrow> \<alpha>\<Colon>group \<Rightarrow> \<alpha>\<Colon>group"} with the corresponding theorem @{thm
+  pow_int_def [no_vars]}.
 *}
 
 subsection {* A note on syntax *}
 
 text {*
-  As a convenience, class context syntax allows references
-  to local class operations and their global counterparts
-  uniformly;  type inference resolves ambiguities.  For example:
+  As a convenience, class context syntax allows references to local
+  class operations and their global counterparts uniformly; type
+  inference resolves ambiguities.  For example:
 *}
 
 context %quote semigroup
@@ -581,11 +572,11 @@
 term %quote "x \<otimes> y" -- {* example 3 *}
 
 text {*
-  \noindent Here in example 1, the term refers to the local class operation
-  @{text "mult [\<alpha>]"}, whereas in example 2 the type constraint
-  enforces the global class operation @{text "mult [nat]"}.
-  In the global context in example 3, the reference is
-  to the polymorphic global class operation @{text "mult [?\<alpha> \<Colon> semigroup]"}.
+  \noindent Here in example 1, the term refers to the local class
+  operation @{text "mult [\<alpha>]"}, whereas in example 2 the type
+  constraint enforces the global class operation @{text "mult [nat]"}.
+  In the global context in example 3, the reference is to the
+  polymorphic global class operation @{text "mult [?\<alpha> \<Colon> semigroup]"}.
 *}
 
 section {* Further issues *}
@@ -593,16 +584,14 @@
 subsection {* Type classes and code generation *}
 
 text {*
-  Turning back to the first motivation for type classes,
-  namely overloading, it is obvious that overloading
-  stemming from @{command class} statements and
-  @{command instantiation}
-  targets naturally maps to Haskell type classes.
-  The code generator framework \cite{isabelle-codegen} 
-  takes this into account.  If the target language (e.g.~SML)
-  lacks type classes, then they
-  are implemented by an explicit dictionary construction.
-  As example, let's go back to the power function:
+  Turning back to the first motivation for type classes, namely
+  overloading, it is obvious that overloading stemming from @{command
+  class} statements and @{command instantiation} targets naturally
+  maps to Haskell type classes.  The code generator framework
+  \cite{isabelle-codegen} takes this into account.  If the target
+  language (e.g.~SML) lacks type classes, then they are implemented by
+  an explicit dictionary construction.  As example, let's go back to
+  the power function:
 *}
 
 definition %quote example :: int where
@@ -619,11 +608,18 @@
 *}
 text %quote {*@{code_stmts example (SML)}*}
 
+text {*
+  \noindent In Scala, implicts are used as dictionaries:
+*}
+(*<*)code_include %invisible Scala "Natural" -(*>*)
+text %quote {*@{code_stmts example (Scala)}*}
+
+
 subsection {* Inspecting the type class universe *}
 
 text {*
-  To facilitate orientation in complex subclass structures,
-  two diagnostics commands are provided:
+  To facilitate orientation in complex subclass structures, two
+  diagnostics commands are provided:
 
   \begin{description}
 
--- a/doc-src/Classes/Thy/document/Classes.tex	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Classes/Thy/document/Classes.tex	Wed Sep 01 17:19:47 2010 +0200
@@ -26,14 +26,14 @@
 Type classes were introduced by Wadler and Blott \cite{wadler89how}
   into the Haskell language to allow for a reasonable implementation
   of overloading\footnote{throughout this tutorial, we are referring
-  to classical Haskell 1.0 type classes, not considering
-  later additions in expressiveness}.
-  As a canonical example, a polymorphic equality function
-  \isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} which is overloaded on different
-  types for \isa{{\isasymalpha}}, which is achieved by splitting introduction
-  of the \isa{eq} function from its overloaded definitions by means
-  of \isa{class} and \isa{instance} declarations:
-  \footnote{syntax here is a kind of isabellized Haskell}
+  to classical Haskell 1.0 type classes, not considering later
+  additions in expressiveness}.  As a canonical example, a polymorphic
+  equality function \isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} which is overloaded on
+  different types for \isa{{\isasymalpha}}, which is achieved by splitting
+  introduction of the \isa{eq} function from its overloaded
+  definitions by means of \isa{class} and \isa{instance}
+  declarations: \footnote{syntax here is a kind of isabellized
+  Haskell}
 
   \begin{quote}
 
@@ -59,14 +59,14 @@
   these annotations are assertions that a particular polymorphic type
   provides definitions for overloaded functions.
 
-  Indeed, type classes not only allow for simple overloading
-  but form a generic calculus, an instance of order-sorted
-  algebra \cite{nipkow-sorts93,Nipkow-Prehofer:1993,Wenzel:1997:TPHOL}.
+  Indeed, type classes not only allow for simple overloading but form
+  a generic calculus, an instance of order-sorted algebra
+  \cite{nipkow-sorts93,Nipkow-Prehofer:1993,Wenzel:1997:TPHOL}.
 
-  From a software engineering point of view, type classes
-  roughly correspond to interfaces in object-oriented languages like Java;
-  so, it is naturally desirable that type classes do not only
-  provide functions (class parameters) but also state specifications
+  From a software engineering point of view, type classes roughly
+  correspond to interfaces in object-oriented languages like Java; so,
+  it is naturally desirable that type classes do not only provide
+  functions (class parameters) but also state specifications
   implementations must obey.  For example, the \isa{class\ eq}
   above could be given the following specification, demanding that
   \isa{class\ eq} is an equivalence relation obeying reflexivity,
@@ -83,11 +83,10 @@
 
   \end{quote}
 
-  \noindent From a theoretical point of view, type classes are lightweight
-  modules; Haskell type classes may be emulated by
-  SML functors \cite{classes_modules}. 
-  Isabelle/Isar offers a discipline of type classes which brings
-  all those aspects together:
+  \noindent From a theoretical point of view, type classes are
+  lightweight modules; Haskell type classes may be emulated by SML
+  functors \cite{classes_modules}.  Isabelle/Isar offers a discipline
+  of type classes which brings all those aspects together:
 
   \begin{enumerate}
     \item specifying abstract parameters together with
@@ -99,15 +98,15 @@
       locales \cite{kammueller-locales}.
   \end{enumerate}
 
-  \noindent Isar type classes also directly support code generation
-  in a Haskell like fashion. Internally, they are mapped to more primitive 
-  Isabelle concepts \cite{Haftmann-Wenzel:2006:classes}.
+  \noindent Isar type classes also directly support code generation in
+  a Haskell like fashion. Internally, they are mapped to more
+  primitive Isabelle concepts \cite{Haftmann-Wenzel:2006:classes}.
 
-  This tutorial demonstrates common elements of structured specifications
-  and abstract reasoning with type classes by the algebraic hierarchy of
-  semigroups, monoids and groups.  Our background theory is that of
-  Isabelle/HOL \cite{isa-tutorial}, for which some
-  familiarity is assumed.%
+  This tutorial demonstrates common elements of structured
+  specifications and abstract reasoning with type classes by the
+  algebraic hierarchy of semigroups, monoids and groups.  Our
+  background theory is that of Isabelle/HOL \cite{isa-tutorial}, for
+  which some familiarity is assumed.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -142,12 +141,9 @@
 \endisadelimquote
 %
 \begin{isamarkuptext}%
-\noindent This \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} specification consists of two
-  parts: the \qn{operational} part names the class parameter
-  (\hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}}), the \qn{logical} part specifies properties on them
-  (\hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}}).  The local \hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}} and
-  \hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}} are lifted to the theory toplevel,
-  yielding the global
+\noindent This \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} specification consists of two parts:
+  the \qn{operational} part names the class parameter (\hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}}), the \qn{logical} part specifies properties on them
+  (\hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}}).  The local \hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}} and \hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}} are lifted to the theory toplevel, yielding the global
   parameter \isa{{\isachardoublequote}mult\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequote}} and the
   global theorem \hyperlink{fact.semigroup.assoc:}{\mbox{\isa{semigroup{\isachardot}assoc{\isacharcolon}}}}~\isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup{\isachardot}\ {\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequote}}.%
 \end{isamarkuptext}%
@@ -158,10 +154,9 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-The concrete type \isa{int} is made a \isa{semigroup}
-  instance by providing a suitable definition for the class parameter
-  \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} and a proof for the specification of \hyperlink{fact.assoc}{\mbox{\isa{assoc}}}.
-  This is accomplished by the \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} target:%
+The concrete type \isa{int} is made a \isa{semigroup} instance
+  by providing a suitable definition for the class parameter \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} and a proof for the specification of \hyperlink{fact.assoc}{\mbox{\isa{assoc}}}.  This is
+  accomplished by the \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} target:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -204,22 +199,19 @@
 \endisadelimquote
 %
 \begin{isamarkuptext}%
-\noindent \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} defines class parameters
-  at a particular instance using common specification tools (here,
-  \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}).  The concluding \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}}
-  opens a proof that the given parameters actually conform
-  to the class specification.  Note that the first proof step
-  is the \hyperlink{method.default}{\mbox{\isa{default}}} method,
-  which for such instance proofs maps to the \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isacharunderscore}classes}}} method.
-  This reduces an instance judgement to the relevant primitive
-  proof goals; typically it is the first method applied
-  in an instantiation proof.
+\noindent \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} defines class parameters at a
+  particular instance using common specification tools (here,
+  \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}).  The concluding \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} opens a
+  proof that the given parameters actually conform to the class
+  specification.  Note that the first proof step is the \hyperlink{method.default}{\mbox{\isa{default}}} method, which for such instance proofs maps to the \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isacharunderscore}classes}}} method.  This reduces an instance judgement to the
+  relevant primitive proof goals; typically it is the first method
+  applied in an instantiation proof.
 
-  From now on, the type-checker will consider \isa{int}
-  as a \isa{semigroup} automatically, i.e.\ any general results
-  are immediately available on concrete instances.
+  From now on, the type-checker will consider \isa{int} as a \isa{semigroup} automatically, i.e.\ any general results are immediately
+  available on concrete instances.
 
-  \medskip Another instance of \isa{semigroup} yields the natural numbers:%
+  \medskip Another instance of \isa{semigroup} yields the natural
+  numbers:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -259,12 +251,12 @@
 \endisadelimquote
 %
 \begin{isamarkuptext}%
-\noindent Note the occurence of the name \isa{mult{\isacharunderscore}nat}
-  in the primrec declaration;  by default, the local name of
-  a class operation \isa{f} to be instantiated on type constructor
-  \isa{{\isasymkappa}} is mangled as \isa{f{\isacharunderscore}{\isasymkappa}}.  In case of uncertainty,
-  these names may be inspected using the \hyperlink{command.print-context}{\mbox{\isa{\isacommand{print{\isacharunderscore}context}}}} command
-  or the corresponding ProofGeneral button.%
+\noindent Note the occurence of the name \isa{mult{\isacharunderscore}nat} in the
+  primrec declaration; by default, the local name of a class operation
+  \isa{f} to be instantiated on type constructor \isa{{\isasymkappa}} is
+  mangled as \isa{f{\isacharunderscore}{\isasymkappa}}.  In case of uncertainty, these names may be
+  inspected using the \hyperlink{command.print-context}{\mbox{\isa{\isacommand{print{\isacharunderscore}context}}}} command or the
+  corresponding ProofGeneral button.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -273,10 +265,9 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-Overloaded definitions given at a class instantiation
-  may include recursion over the syntactic structure of types.
-  As a canonical example, we model product semigroups
-  using our simple algebra:%
+Overloaded definitions given at a class instantiation may include
+  recursion over the syntactic structure of types.  As a canonical
+  example, we model product semigroups using our simple algebra:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -318,11 +309,10 @@
 \begin{isamarkuptext}%
 \noindent Associativity of product semigroups is established using
   the definition of \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} on products and the hypothetical
-  associativity of the type components;  these hypotheses
-  are legitimate due to the \isa{semigroup} constraints imposed
-  on the type components by the \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} proposition.
-  Indeed, this pattern often occurs with parametric types
-  and type classes.%
+  associativity of the type components; these hypotheses are
+  legitimate due to the \isa{semigroup} constraints imposed on the
+  type components by the \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} proposition.  Indeed,
+  this pattern often occurs with parametric types and type classes.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -331,10 +321,9 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-We define a subclass \isa{monoidl} (a semigroup with a left-hand neutral)
-  by extending \isa{semigroup}
-  with one additional parameter \isa{neutral} together
-  with its characteristic property:%
+We define a subclass \isa{monoidl} (a semigroup with a left-hand
+  neutral) by extending \isa{semigroup} with one additional
+  parameter \isa{neutral} together with its characteristic property:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -355,10 +344,10 @@
 \endisadelimquote
 %
 \begin{isamarkuptext}%
-\noindent Again, we prove some instances, by
-  providing suitable parameter definitions and proofs for the
-  additional specifications.  Observe that instantiations
-  for types with the same arity may be simultaneous:%
+\noindent Again, we prove some instances, by providing suitable
+  parameter definitions and proofs for the additional specifications.
+  Observe that instantiations for types with the same arity may be
+  simultaneous:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -505,8 +494,7 @@
 \endisadelimquote
 %
 \begin{isamarkuptext}%
-\noindent To finish our small algebra example, we add a \isa{group} class
-  with a corresponding instance:%
+\noindent To finish our small algebra example, we add a \isa{group} class with a corresponding instance:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -563,9 +551,9 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-The example above gives an impression how Isar type classes work
-  in practice.  As stated in the introduction, classes also provide
-  a link to Isar's locale system.  Indeed, the logical core of a class
+The example above gives an impression how Isar type classes work in
+  practice.  As stated in the introduction, classes also provide a
+  link to Isar's locale system.  Indeed, the logical core of a class
   is nothing other than a locale:%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -780,10 +768,12 @@
 \endisadelimquote
 %
 \begin{isamarkuptext}%
-\noindent Here the \qt{\hyperlink{keyword.in}{\mbox{\isa{\isakeyword{in}}}} \isa{group}} target specification
-  indicates that the result is recorded within that context for later
-  use.  This local theorem is also lifted to the global one \hyperlink{fact.group.left-cancel:}{\mbox{\isa{group{\isachardot}left{\isacharunderscore}cancel{\isacharcolon}}}} \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}group{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}.  Since type \isa{int} has been made an instance of
-  \isa{group} before, we may refer to that fact as well: \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ int{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}.%
+\noindent Here the \qt{\hyperlink{keyword.in}{\mbox{\isa{\isakeyword{in}}}} \isa{group}} target
+  specification indicates that the result is recorded within that
+  context for later use.  This local theorem is also lifted to the
+  global one \hyperlink{fact.group.left-cancel:}{\mbox{\isa{group{\isachardot}left{\isacharunderscore}cancel{\isacharcolon}}}} \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}group{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}.  Since type \isa{int} has been
+  made an instance of \isa{group} before, we may refer to that
+  fact as well: \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ int{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -814,16 +804,14 @@
 %
 \begin{isamarkuptext}%
 \noindent If the locale \isa{group} is also a class, this local
-  definition is propagated onto a global definition of
-  \isa{{\isachardoublequote}pow{\isacharunderscore}nat\ {\isasymColon}\ nat\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid{\isachardoublequote}}
-  with corresponding theorems
+  definition is propagated onto a global definition of \isa{{\isachardoublequote}pow{\isacharunderscore}nat\ {\isasymColon}\ nat\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid{\isachardoublequote}} with corresponding theorems
 
   \isa{pow{\isacharunderscore}nat\ {\isadigit{0}}\ x\ {\isacharequal}\ {\isasymone}\isasep\isanewline%
 pow{\isacharunderscore}nat\ {\isacharparenleft}Suc\ n{\isacharparenright}\ x\ {\isacharequal}\ x\ {\isasymotimes}\ pow{\isacharunderscore}nat\ n\ x}.
 
-  \noindent As you can see from this example, for local
-  definitions you may use any specification tool
-  which works together with locales, such as Krauss's recursive function package
+  \noindent As you can see from this example, for local definitions
+  you may use any specification tool which works together with
+  locales, such as Krauss's recursive function package
   \cite{krauss2006}.%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -833,19 +821,17 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-We introduced Isar classes by analogy to type classes in
-  functional programming;  if we reconsider this in the
-  context of what has been said about type classes and locales,
-  we can drive this analogy further by stating that type
-  classes essentially correspond to functors that have
-  a canonical interpretation as type classes.
-  There is also the possibility of other interpretations.
-  For example, \isa{list}s also form a monoid with
-  \isa{append} and \isa{{\isacharbrackleft}{\isacharbrackright}} as operations, but it
-  seems inappropriate to apply to lists
-  the same operations as for genuinely algebraic types.
-  In such a case, we can simply make a particular interpretation
-  of monoids for lists:%
+We introduced Isar classes by analogy to type classes in functional
+  programming; if we reconsider this in the context of what has been
+  said about type classes and locales, we can drive this analogy
+  further by stating that type classes essentially correspond to
+  functors that have a canonical interpretation as type classes.
+  There is also the possibility of other interpretations.  For
+  example, \isa{list}s also form a monoid with \isa{append} and
+  \isa{{\isacharbrackleft}{\isacharbrackright}} as operations, but it seems inappropriate to apply to
+  lists the same operations as for genuinely algebraic types.  In such
+  a case, we can simply make a particular interpretation of monoids
+  for lists:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -969,12 +955,10 @@
 \endisadelimquote
 %
 \begin{isamarkuptext}%
-The logical proof is carried out on the locale level.
-  Afterwards it is propagated
-  to the type system, making \isa{group} an instance of
-  \isa{monoid} by adding an additional edge
-  to the graph of subclass relations
-  (\figref{fig:subclass}).
+The logical proof is carried out on the locale level.  Afterwards it
+  is propagated to the type system, making \isa{group} an instance
+  of \isa{monoid} by adding an additional edge to the graph of
+  subclass relations (\figref{fig:subclass}).
 
   \begin{figure}[htbp]
    \begin{center}
@@ -1006,8 +990,7 @@
    \end{center}
   \end{figure}
 
-  For illustration, a derived definition
-  in \isa{group} using \isa{pow{\isacharunderscore}nat}%
+  For illustration, a derived definition in \isa{group} using \isa{pow{\isacharunderscore}nat}%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1029,9 +1012,7 @@
 \endisadelimquote
 %
 \begin{isamarkuptext}%
-\noindent yields the global definition of
-  \isa{{\isachardoublequote}pow{\isacharunderscore}int\ {\isasymColon}\ int\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group{\isachardoublequote}}
-  with the corresponding theorem \isa{pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ {\isadigit{0}}\ {\isasymle}\ k\ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}}.%
+\noindent yields the global definition of \isa{{\isachardoublequote}pow{\isacharunderscore}int\ {\isasymColon}\ int\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group{\isachardoublequote}} with the corresponding theorem \isa{pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ {\isadigit{0}}\ {\isasymle}\ k\ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}}.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1040,9 +1021,9 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-As a convenience, class context syntax allows references
-  to local class operations and their global counterparts
-  uniformly;  type inference resolves ambiguities.  For example:%
+As a convenience, class context syntax allows references to local
+  class operations and their global counterparts uniformly; type
+  inference resolves ambiguities.  For example:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1082,11 +1063,11 @@
 \endisadelimquote
 %
 \begin{isamarkuptext}%
-\noindent Here in example 1, the term refers to the local class operation
-  \isa{mult\ {\isacharbrackleft}{\isasymalpha}{\isacharbrackright}}, whereas in example 2 the type constraint
-  enforces the global class operation \isa{mult\ {\isacharbrackleft}nat{\isacharbrackright}}.
-  In the global context in example 3, the reference is
-  to the polymorphic global class operation \isa{mult\ {\isacharbrackleft}{\isacharquery}{\isasymalpha}\ {\isasymColon}\ semigroup{\isacharbrackright}}.%
+\noindent Here in example 1, the term refers to the local class
+  operation \isa{mult\ {\isacharbrackleft}{\isasymalpha}{\isacharbrackright}}, whereas in example 2 the type
+  constraint enforces the global class operation \isa{mult\ {\isacharbrackleft}nat{\isacharbrackright}}.
+  In the global context in example 3, the reference is to the
+  polymorphic global class operation \isa{mult\ {\isacharbrackleft}{\isacharquery}{\isasymalpha}\ {\isasymColon}\ semigroup{\isacharbrackright}}.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1099,16 +1080,13 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-Turning back to the first motivation for type classes,
-  namely overloading, it is obvious that overloading
-  stemming from \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} statements and
-  \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}}
-  targets naturally maps to Haskell type classes.
-  The code generator framework \cite{isabelle-codegen} 
-  takes this into account.  If the target language (e.g.~SML)
-  lacks type classes, then they
-  are implemented by an explicit dictionary construction.
-  As example, let's go back to the power function:%
+Turning back to the first motivation for type classes, namely
+  overloading, it is obvious that overloading stemming from \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} statements and \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} targets naturally
+  maps to Haskell type classes.  The code generator framework
+  \cite{isabelle-codegen} takes this into account.  If the target
+  language (e.g.~SML) lacks type classes, then they are implemented by
+  an explicit dictionary construction.  As example, let's go back to
+  the power function:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1328,13 +1306,121 @@
 %
 \endisadelimquote
 %
+\begin{isamarkuptext}%
+\noindent In Scala, implicts are used as dictionaries:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadeliminvisible
+%
+\endisadeliminvisible
+%
+\isataginvisible
+%
+\endisataginvisible
+{\isafoldinvisible}%
+%
+\isadeliminvisible
+%
+\endisadeliminvisible
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\isatagquote
+%
+\begin{isamarkuptext}%
+\isatypewriter%
+\noindent%
+\hspace*{0pt}object Example {\char123}\\
+\hspace*{0pt}\\
+\hspace*{0pt}import /*implicits*/ Example.semigroup{\char95}int,~Example.monoidl{\char95}int,\\
+\hspace*{0pt} ~Example.monoid{\char95}int,~Example.group{\char95}int\\
+\hspace*{0pt}\\
+\hspace*{0pt}abstract sealed class nat\\
+\hspace*{0pt}final case object Zero{\char95}nat extends nat\\
+\hspace*{0pt}final case class Suc(a:~nat) extends nat\\
+\hspace*{0pt}\\
+\hspace*{0pt}def nat{\char95}aux(i:~BigInt,~n:~nat):~nat =\\
+\hspace*{0pt} ~(if (i <= BigInt(0)) n else nat{\char95}aux(i - BigInt(1),~Suc(n)))\\
+\hspace*{0pt}\\
+\hspace*{0pt}def nat(i:~BigInt):~nat = nat{\char95}aux(i,~Zero{\char95}nat)\\
+\hspace*{0pt}\\
+\hspace*{0pt}trait semigroup[A] {\char123}\\
+\hspace*{0pt} ~val `Example+mult`:~(A,~A) => A\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}def mult[A](a:~A,~b:~A)(implicit A:~semigroup[A]):~A =\\
+\hspace*{0pt} ~A.`Example+mult`(a,~b)\\
+\hspace*{0pt}\\
+\hspace*{0pt}trait monoidl[A] extends semigroup[A] {\char123}\\
+\hspace*{0pt} ~val `Example+neutral`:~A\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}def neutral[A](implicit A:~monoidl[A]):~A = A.`Example+neutral`\\
+\hspace*{0pt}\\
+\hspace*{0pt}trait monoid[A] extends monoidl[A] {\char123}\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}\\
+\hspace*{0pt}trait group[A] extends monoid[A] {\char123}\\
+\hspace*{0pt} ~val `Example+inverse`:~A => A\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}def inverse[A](a:~A)(implicit A:~group[A]):~A = A.`Example+inverse`(a)\\
+\hspace*{0pt}\\
+\hspace*{0pt}def pow{\char95}nat[A:~monoid](xa0:~nat,~x:~A):~A = (xa0,~x) match {\char123}\\
+\hspace*{0pt} ~case (Zero{\char95}nat,~x) => neutral[A]\\
+\hspace*{0pt} ~case (Suc(n),~x) => mult[A](x,~pow{\char95}nat[A](n,~x))\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}\\
+\hspace*{0pt}def pow{\char95}int[A:~group](k:~BigInt,~x:~A):~A =\\
+\hspace*{0pt} ~(if (BigInt(0) <= k) pow{\char95}nat[A](nat(k),~x)\\
+\hspace*{0pt} ~~~else inverse[A](pow{\char95}nat[A](nat((- k)),~x)))\\
+\hspace*{0pt}\\
+\hspace*{0pt}def mult{\char95}int(i:~BigInt,~j:~BigInt):~BigInt = i + j\\
+\hspace*{0pt}\\
+\hspace*{0pt}implicit def semigroup{\char95}int:~semigroup[BigInt] = new semigroup[BigInt] {\char123}\\
+\hspace*{0pt} ~val `Example+mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}\\
+\hspace*{0pt}def neutral{\char95}int:~BigInt = BigInt(0)\\
+\hspace*{0pt}\\
+\hspace*{0pt}implicit def monoidl{\char95}int:~monoidl[BigInt] = new monoidl[BigInt] {\char123}\\
+\hspace*{0pt} ~val `Example+neutral` = neutral{\char95}int\\
+\hspace*{0pt} ~val `Example+mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}\\
+\hspace*{0pt}implicit def monoid{\char95}int:~monoid[BigInt] = new monoid[BigInt] {\char123}\\
+\hspace*{0pt} ~val `Example+neutral` = neutral{\char95}int\\
+\hspace*{0pt} ~val `Example+mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}\\
+\hspace*{0pt}def inverse{\char95}int(i:~BigInt):~BigInt = (- i)\\
+\hspace*{0pt}\\
+\hspace*{0pt}implicit def group{\char95}int:~group[BigInt] = new group[BigInt] {\char123}\\
+\hspace*{0pt} ~val `Example+inverse` = (a:~BigInt) => inverse{\char95}int(a)\\
+\hspace*{0pt} ~val `Example+neutral` = neutral{\char95}int\\
+\hspace*{0pt} ~val `Example+mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}\\
+\hspace*{0pt}def example:~BigInt = pow{\char95}int[BigInt](BigInt(10),~BigInt(- 2))\\
+\hspace*{0pt}\\
+\hspace*{0pt}{\char125}~/* object Example */%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagquote
+{\isafoldquote}%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
 \isamarkupsubsection{Inspecting the type class universe%
 }
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-To facilitate orientation in complex subclass structures,
-  two diagnostics commands are provided:
+To facilitate orientation in complex subclass structures, two
+  diagnostics commands are provided:
 
   \begin{description}
 
--- a/doc-src/Codegen/Thy/Foundations.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/Foundations.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -220,12 +220,12 @@
 text {*
   \noindent Obviously, polymorphic equality is implemented the Haskell
   way using a type class.  How is this achieved?  HOL introduces an
-  explicit class @{class eq} with a corresponding operation @{const
-  eq_class.eq} such that @{thm eq [no_vars]}.  The preprocessing
-  framework does the rest by propagating the @{class eq} constraints
+  explicit class @{class equal} with a corresponding operation @{const
+  HOL.equal} such that @{thm equal [no_vars]}.  The preprocessing
+  framework does the rest by propagating the @{class equal} constraints
   through all dependent code equations.  For datatypes, instances of
-  @{class eq} are implicitly derived when possible.  For other types,
-  you may instantiate @{text eq} manually like any other type class.
+  @{class equal} are implicitly derived when possible.  For other types,
+  you may instantiate @{text equal} manually like any other type class.
 *}
 
 
--- a/doc-src/Codegen/Thy/Inductive_Predicate.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/Inductive_Predicate.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -7,7 +7,7 @@
 
 inductive %invisible append where
   "append [] ys ys"
-| "append xs ys zs ==> append (x # xs) ys (x # zs)"
+| "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
 
 lemma %invisible append: "append xs ys zs = (xs @ ys = zs)"
   by (induct xs arbitrary: ys zs) (auto elim: append.cases intro: append.intros)
@@ -95,9 +95,9 @@
   "append_i_i_o"}).  You can specify your own names as follows:
 *}
 
-code_pred %quote (modes: i => i => o => bool as concat,
-  o => o => i => bool as split,
-  i => o => i => bool as suffix) append .
+code_pred %quote (modes: i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool as concat,
+  o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as split,
+  i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as suffix) append .
 
 subsection {* Alternative introduction rules *}
 
@@ -220,8 +220,8 @@
   "values"} and the number of elements.
 *}
 
-values %quote [mode: i => o => bool] 20 "{n. tranclp succ 10 n}"
-values %quote [mode: o => i => bool] 10 "{n. tranclp succ n 10}"
+values %quote [mode: i \<Rightarrow> o \<Rightarrow> bool] 20 "{n. tranclp succ 10 n}"
+values %quote [mode: o \<Rightarrow> i \<Rightarrow> bool] 10 "{n. tranclp succ n 10}"
 
 
 subsection {* Embedding into functional code within Isabelle/HOL *}
--- a/doc-src/Codegen/Thy/Introduction.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/Introduction.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -8,8 +8,9 @@
   This tutorial introduces the code generator facilities of @{text
   "Isabelle/HOL"}.  It allows to turn (a certain class of) HOL
   specifications into corresponding executable code in the programming
-  languages @{text SML} \cite{SML}, @{text OCaml} \cite{OCaml} and
-  @{text Haskell} \cite{haskell-revised-report}.
+  languages @{text SML} \cite{SML}, @{text OCaml} \cite{OCaml},
+  @{text Haskell} \cite{haskell-revised-report} and @{text Scala}
+  \cite{scala-overview-tech-report}.
 
   To profit from this tutorial, some familiarity and experience with
   @{theory HOL} \cite{isa-tutorial} and its basic theories is assumed.
@@ -78,8 +79,8 @@
   target language identifier and a freely chosen module name.  A file
   name denotes the destination to store the generated code.  Note that
   the semantics of the destination depends on the target language: for
-  @{text SML} and @{text OCaml} it denotes a \emph{file}, for @{text
-  Haskell} it denotes a \emph{directory} where a file named as the
+  @{text SML}, @{text OCaml} and @{text Scala} it denotes a \emph{file},
+  for @{text Haskell} it denotes a \emph{directory} where a file named as the
   module name (with extension @{text ".hs"}) is written:
 *}
 
--- a/doc-src/Codegen/Thy/Setup.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/Setup.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -27,6 +27,6 @@
 
 setup {* Code_Target.set_default_code_width 74 *}
 
-ML_command {* Unsynchronized.reset unique_names *}
+ML_command {* unique_names := false *}
 
 end
--- a/doc-src/Codegen/Thy/document/Adaptation.tex	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Adaptation.tex	Wed Sep 01 17:19:47 2010 +0200
@@ -240,7 +240,7 @@
 \hspace*{0pt}structure Example :~sig\\
 \hspace*{0pt} ~datatype nat = Zero{\char95}nat | Suc of nat\\
 \hspace*{0pt} ~datatype boola = True | False\\
-\hspace*{0pt} ~val anda :~boola -> boola -> boola\\
+\hspace*{0pt} ~val conj :~boola -> boola -> boola\\
 \hspace*{0pt} ~val less{\char95}nat :~nat -> nat -> boola\\
 \hspace*{0pt} ~val less{\char95}eq{\char95}nat :~nat -> nat -> boola\\
 \hspace*{0pt} ~val in{\char95}interval :~nat * nat -> nat -> boola\\
@@ -250,17 +250,17 @@
 \hspace*{0pt}\\
 \hspace*{0pt}datatype boola = True | False;\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun anda p True = p\\
-\hspace*{0pt} ~| anda p False = False\\
-\hspace*{0pt} ~| anda True p = p\\
-\hspace*{0pt} ~| anda False p = False;\\
+\hspace*{0pt}fun conj p True = p\\
+\hspace*{0pt} ~| conj p False = False\\
+\hspace*{0pt} ~| conj True p = p\\
+\hspace*{0pt} ~| conj False p = False;\\
 \hspace*{0pt}\\
 \hspace*{0pt}fun less{\char95}nat m (Suc n) = less{\char95}eq{\char95}nat m n\\
 \hspace*{0pt} ~| less{\char95}nat n Zero{\char95}nat = False\\
 \hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\
 \hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = True;\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun in{\char95}interval (k,~l) n = anda (less{\char95}eq{\char95}nat k n) (less{\char95}eq{\char95}nat n l);\\
+\hspace*{0pt}fun in{\char95}interval (k,~l) n = conj (less{\char95}eq{\char95}nat k n) (less{\char95}eq{\char95}nat n l);\\
 \hspace*{0pt}\\
 \hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
--- a/doc-src/Codegen/Thy/document/Inductive_Predicate.tex	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Inductive_Predicate.tex	Wed Sep 01 17:19:47 2010 +0200
@@ -212,9 +212,9 @@
 %
 \isatagquote
 \isacommand{code{\isacharunderscore}pred}\isamarkupfalse%
-\ {\isacharparenleft}modes{\isacharcolon}\ i\ {\isacharequal}{\isachargreater}\ i\ {\isacharequal}{\isachargreater}\ o\ {\isacharequal}{\isachargreater}\ bool\ as\ concat{\isacharcomma}\isanewline
-\ \ o\ {\isacharequal}{\isachargreater}\ o\ {\isacharequal}{\isachargreater}\ i\ {\isacharequal}{\isachargreater}\ bool\ as\ split{\isacharcomma}\isanewline
-\ \ i\ {\isacharequal}{\isachargreater}\ o\ {\isacharequal}{\isachargreater}\ i\ {\isacharequal}{\isachargreater}\ bool\ as\ suffix{\isacharparenright}\ append\ \isacommand{{\isachardot}}\isamarkupfalse%
+\ {\isacharparenleft}modes{\isacharcolon}\ i\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ bool\ as\ concat{\isacharcomma}\isanewline
+\ \ o\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ bool\ as\ split{\isacharcomma}\isanewline
+\ \ i\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ bool\ as\ suffix{\isacharparenright}\ append\ \isacommand{{\isachardot}}\isamarkupfalse%
 %
 \endisatagquote
 {\isafoldquote}%
@@ -422,9 +422,9 @@
 %
 \isatagquote
 \isacommand{values}\isamarkupfalse%
-\ {\isacharbrackleft}mode{\isacharcolon}\ i\ {\isacharequal}{\isachargreater}\ o\ {\isacharequal}{\isachargreater}\ bool{\isacharbrackright}\ {\isadigit{2}}{\isadigit{0}}\ {\isachardoublequoteopen}{\isacharbraceleft}n{\isachardot}\ tranclp\ succ\ {\isadigit{1}}{\isadigit{0}}\ n{\isacharbraceright}{\isachardoublequoteclose}\isanewline
+\ {\isacharbrackleft}mode{\isacharcolon}\ i\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ bool{\isacharbrackright}\ {\isadigit{2}}{\isadigit{0}}\ {\isachardoublequoteopen}{\isacharbraceleft}n{\isachardot}\ tranclp\ succ\ {\isadigit{1}}{\isadigit{0}}\ n{\isacharbraceright}{\isachardoublequoteclose}\isanewline
 \isacommand{values}\isamarkupfalse%
-\ {\isacharbrackleft}mode{\isacharcolon}\ o\ {\isacharequal}{\isachargreater}\ i\ {\isacharequal}{\isachargreater}\ bool{\isacharbrackright}\ {\isadigit{1}}{\isadigit{0}}\ {\isachardoublequoteopen}{\isacharbraceleft}n{\isachardot}\ tranclp\ succ\ n\ {\isadigit{1}}{\isadigit{0}}{\isacharbraceright}{\isachardoublequoteclose}%
+\ {\isacharbrackleft}mode{\isacharcolon}\ o\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ bool{\isacharbrackright}\ {\isadigit{1}}{\isadigit{0}}\ {\isachardoublequoteopen}{\isacharbraceleft}n{\isachardot}\ tranclp\ succ\ n\ {\isadigit{1}}{\isadigit{0}}{\isacharbraceright}{\isachardoublequoteclose}%
 \endisatagquote
 {\isafoldquote}%
 %
--- a/doc-src/Codegen/Thy/document/Introduction.tex	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Introduction.tex	Wed Sep 01 17:19:47 2010 +0200
@@ -25,8 +25,9 @@
 \begin{isamarkuptext}%
 This tutorial introduces the code generator facilities of \isa{Isabelle{\isacharslash}HOL}.  It allows to turn (a certain class of) HOL
   specifications into corresponding executable code in the programming
-  languages \isa{SML} \cite{SML}, \isa{OCaml} \cite{OCaml} and
-  \isa{Haskell} \cite{haskell-revised-report}.
+  languages \isa{SML} \cite{SML}, \isa{OCaml} \cite{OCaml},
+  \isa{Haskell} \cite{haskell-revised-report} and \isa{Scala}
+  \cite{scala-overview-tech-report}.
 
   To profit from this tutorial, some familiarity and experience with
   \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} \cite{isa-tutorial} and its basic theories is assumed.%
@@ -191,7 +192,8 @@
   target language identifier and a freely chosen module name.  A file
   name denotes the destination to store the generated code.  Note that
   the semantics of the destination depends on the target language: for
-  \isa{SML} and \isa{OCaml} it denotes a \emph{file}, for \isa{Haskell} it denotes a \emph{directory} where a file named as the
+  \isa{SML}, \isa{OCaml} and \isa{Scala} it denotes a \emph{file},
+  for \isa{Haskell} it denotes a \emph{directory} where a file named as the
   module name (with extension \isa{{\isachardot}hs}) is written:%
 \end{isamarkuptext}%
 \isamarkuptrue%
--- a/doc-src/Codegen/Thy/pictures/architecture.tex	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/pictures/architecture.tex	Wed Sep 01 17:19:47 2010 +0200
@@ -30,8 +30,8 @@
   \node (seri) at (1.5, 0) [process, positive] {serialisation};
   \node (SML) at (2.5, 3) [entity, positive] {\sys{SML}};
   \node (OCaml) at (2.5, 2) [entity, positive] {\sys{OCaml}};
-  \node (further) at (2.5, 1) [entity, positive] {(\ldots)};
-  \node (Haskell) at (2.5, 0) [entity, positive] {\sys{Haskell}};
+  \node (Haskell) at (2.5, 1) [entity, positive] {\sys{Haskell}};
+  \node (Scala) at (2.5, 0) [entity, positive] {\sys{Scala}};
   \draw [semithick] (spec) -- (spec_user_join);
   \draw [semithick] (user) -- (spec_user_join);
   \draw [-diamond, semithick] (spec_user_join) -- (raw);
@@ -41,8 +41,8 @@
   \draw [arrow] (iml) -- (seri);
   \draw [arrow] (seri) -- (SML);
   \draw [arrow] (seri) -- (OCaml);
-  \draw [arrow, dashed] (seri) -- (further);
   \draw [arrow] (seri) -- (Haskell);
+  \draw [arrow] (seri) -- (Scala);
 \end{tikzpicture}
 
 }
--- a/doc-src/Codegen/codegen.tex	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/codegen.tex	Wed Sep 01 17:19:47 2010 +0200
@@ -22,7 +22,7 @@
 \begin{abstract}
   \noindent This tutorial introduces the code generator facilities of Isabelle/HOL.
     They empower the user to turn HOL specifications into corresponding executable
-    programs in the languages SML, OCaml and Haskell.
+    programs in the languages SML, OCaml, Haskell and Scala.
 \end{abstract}
 
 \thispagestyle{empty}\clearpage
--- a/doc-src/IsarOverview/Isar/ROOT.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/IsarOverview/Isar/ROOT.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -1,3 +1,3 @@
-Unsynchronized.set quick_and_dirty;
+quick_and_dirty := true;
 
 use_thys ["Logic", "Induction"];
--- a/doc-src/IsarRef/Thy/HOL_Specific.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/IsarRef/Thy/HOL_Specific.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -968,7 +968,8 @@
 
   \medskip One framework generates code from functional programs
   (including overloading using type classes) to SML \cite{SML}, OCaml
-  \cite{OCaml} and Haskell \cite{haskell-revised-report}.
+  \cite{OCaml}, Haskell \cite{haskell-revised-report} and Scala
+  \cite{scala-overview-tech-report}.
   Conceptually, code generation is split up in three steps:
   \emph{selection} of code theorems, \emph{translation} into an
   abstract executable view and \emph{serialization} to a specific
@@ -1015,7 +1016,7 @@
     class: nameref
     ;
 
-    target: 'OCaml' | 'SML' | 'Haskell'
+    target: 'SML' | 'OCaml' | 'Haskell' | 'Scala'
     ;
 
     'code' ( 'del' | 'abstype' | 'abstract' ) ?
@@ -1088,7 +1089,7 @@
   after the @{keyword "module_name"} keyword; then \emph{all} code is
   placed in this module.
 
-  For \emph{SML} and \emph{OCaml}, the file specification refers to a
+  For \emph{SML}, \emph{OCaml} and \emph{Scala} the file specification refers to a
   single file; for \emph{Haskell}, it refers to a whole directory,
   where code is generated in multiple files reflecting the module
   hierarchy.  Omitting the file specification denotes standard
--- a/doc-src/IsarRef/Thy/ROOT-HOLCF.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/IsarRef/Thy/ROOT-HOLCF.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -1,4 +1,4 @@
-Unsynchronized.set Thy_Output.source;
+Thy_Output.source_default := true;
 use "../../antiquote_setup.ML";
 
 use_thy "HOLCF_Specific";
--- a/doc-src/IsarRef/Thy/ROOT-ZF.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/IsarRef/Thy/ROOT-ZF.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -1,4 +1,4 @@
-Unsynchronized.set Thy_Output.source;
+Thy_Output.source_default := true;
 use "../../antiquote_setup.ML";
 
 use_thy "ZF_Specific";
--- a/doc-src/IsarRef/Thy/ROOT.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/IsarRef/Thy/ROOT.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -1,5 +1,5 @@
-Unsynchronized.set quick_and_dirty;
-Unsynchronized.set Thy_Output.source;
+quick_and_dirty := true;
+Thy_Output.source_default := true;
 use "../../antiquote_setup.ML";
 
 use_thys [
--- a/doc-src/IsarRef/Thy/document/HOL_Specific.tex	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/IsarRef/Thy/document/HOL_Specific.tex	Wed Sep 01 17:19:47 2010 +0200
@@ -984,7 +984,8 @@
 
   \medskip One framework generates code from functional programs
   (including overloading using type classes) to SML \cite{SML}, OCaml
-  \cite{OCaml} and Haskell \cite{haskell-revised-report}.
+  \cite{OCaml}, Haskell \cite{haskell-revised-report} and Scala
+  \cite{scala-overview-tech-report}.
   Conceptually, code generation is split up in three steps:
   \emph{selection} of code theorems, \emph{translation} into an
   abstract executable view and \emph{serialization} to a specific
@@ -1031,7 +1032,7 @@
     class: nameref
     ;
 
-    target: 'OCaml' | 'SML' | 'Haskell'
+    target: 'SML' | 'OCaml' | 'Haskell' | 'Scala'
     ;
 
     'code' ( 'del' | 'abstype' | 'abstract' ) ?
@@ -1103,7 +1104,7 @@
   after the \hyperlink{keyword.module-name}{\mbox{\isa{\isakeyword{module{\isacharunderscore}name}}}} keyword; then \emph{all} code is
   placed in this module.
 
-  For \emph{SML} and \emph{OCaml}, the file specification refers to a
+  For \emph{SML}, \emph{OCaml} and \emph{Scala} the file specification refers to a
   single file; for \emph{Haskell}, it refers to a whole directory,
   where code is generated in multiple files reflecting the module
   hierarchy.  Omitting the file specification denotes standard
--- a/doc-src/LaTeXsugar/Sugar/Sugar.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/LaTeXsugar/Sugar/Sugar.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -132,7 +132,7 @@
 This \verb!no_vars! business can become a bit tedious.
 If you would rather never see question marks, simply put
 \begin{quote}
-@{ML "Unsynchronized.reset show_question_marks"}\verb!;!
+@{ML "show_question_marks := false"}\verb!;!
 \end{quote}
 at the beginning of your file \texttt{ROOT.ML}.
 The rest of this document is produced with this flag set to \texttt{false}.
@@ -144,7 +144,7 @@
 turning the last digit into a subscript: write \verb!x\<^isub>1! and
 obtain the much nicer @{text"x\<^isub>1"}. *}
 
-(*<*)ML "Unsynchronized.reset show_question_marks"(*>*)
+(*<*)ML "show_question_marks := false"(*>*)
 
 subsection {*Qualified names*}
 
@@ -155,7 +155,7 @@
 short names (no qualifiers) by setting \verb!short_names!, typically
 in \texttt{ROOT.ML}:
 \begin{quote}
-@{ML "Unsynchronized.set short_names"}\verb!;!
+@{ML "short_names := true"}\verb!;!
 \end{quote}
 *}
 
--- a/doc-src/LaTeXsugar/Sugar/document/Sugar.tex	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/LaTeXsugar/Sugar/document/Sugar.tex	Wed Sep 01 17:19:47 2010 +0200
@@ -173,7 +173,7 @@
 This \verb!no_vars! business can become a bit tedious.
 If you would rather never see question marks, simply put
 \begin{quote}
-\verb|Unsynchronized.reset show_question_marks|\verb!;!
+\verb|show_question_marks := false|\verb!;!
 \end{quote}
 at the beginning of your file \texttt{ROOT.ML}.
 The rest of this document is produced with this flag set to \texttt{false}.
@@ -211,7 +211,7 @@
 short names (no qualifiers) by setting \verb!short_names!, typically
 in \texttt{ROOT.ML}:
 \begin{quote}
-\verb|Unsynchronized.set short_names|\verb!;!
+\verb|short_names := true|\verb!;!
 \end{quote}%
 \end{isamarkuptext}%
 \isamarkuptrue%
--- a/doc-src/Main/Docs/Main_Doc.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Main/Docs/Main_Doc.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -10,9 +10,9 @@
    Syntax.pretty_typ ctxt T)
 
 val _ = Thy_Output.antiquotation "term_type_only" (Args.term -- Args.typ_abbrev)
-  (fn {source, context, ...} => fn arg =>
-    Thy_Output.output
-      (Thy_Output.maybe_pretty_source (pretty_term_type_only context) source [arg]));
+  (fn {source, context = ctxt, ...} => fn arg =>
+    Thy_Output.output ctxt
+      (Thy_Output.maybe_pretty_source pretty_term_type_only ctxt source [arg]));
 *}
 (*>*)
 text{*
--- a/doc-src/Sledgehammer/sledgehammer.tex	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Sledgehammer/sledgehammer.tex	Wed Sep 01 17:19:47 2010 +0200
@@ -447,33 +447,29 @@
 \label{relevance-filter}
 
 \begin{enum}
-\opdefault{relevance\_threshold}{int}{40}
-Specifies the threshold above which facts are considered relevant by the
-relevance filter. The option ranges from 0 to 100, where 0 means that all
-theorems are relevant.
+\opdefault{relevance\_thresholds}{int\_pair}{45~85}
+Specifies the thresholds above which facts are considered relevant by the
+relevance filter. The first threshold is used for the first iteration of the
+relevance filter and the second threshold is used for the last iteration (if it
+is reached). The effective threshold is quadratically interpolated for the other
+iterations. Each threshold ranges from 0 to 100, where 0 means that all theorems
+are relevant and 100 only theorems that refer to previously seen constants.
 
-\opdefault{relevance\_convergence}{int}{31}
-Specifies the convergence factor, expressed as a percentage, used by the
-relevance filter. This factor is used by the relevance filter to scale down the
-relevance of facts at each iteration of the filter.
-
-\opdefault{max\_relevant\_per\_iter}{int\_or\_smart}{\textit{smart}}
-Specifies the maximum number of facts that may be added during one iteration of
-the relevance filter. If the option is set to \textit{smart}, it is set to a
-value that was empirically found to be appropriate for the ATP. A typical value
-would be 50.
+\opdefault{max\_relevant}{int\_or\_smart}{\textit{smart}}
+Specifies the maximum number of facts that may be returned by the relevance
+filter. If the option is set to \textit{smart}, it is set to a value that was
+empirically found to be appropriate for the ATP. A typical value would be 300.
 
 \opsmartx{theory\_relevant}{theory\_irrelevant}
 Specifies whether the theory from which a fact comes should be taken into
 consideration by the relevance filter. If the option is set to \textit{smart},
-it is taken to be \textit{true} for SPASS and \textit{false} for E and Vampire,
-because empirical results suggest that these are the best settings.
+it is taken to be \textit{true} for SPASS and \textit{false} for the other ATPs;
+empirical results suggest that these are the best settings.
 
-\opfalse{defs\_relevant}{defs\_irrelevant}
-Specifies whether the definition of constants occurring in the formula to prove
-should be considered particularly relevant. Enabling this option tends to lead
-to larger problems and typically slows down the ATPs.
-
+%\opfalse{defs\_relevant}{defs\_irrelevant}
+%Specifies whether the definition of constants occurring in the formula to prove
+%should be considered particularly relevant. Enabling this option tends to lead
+%to larger problems and typically slows down the ATPs.
 \end{enum}
 
 \subsection{Output Format}
--- a/doc-src/System/Thy/ROOT.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/System/Thy/ROOT.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -1,4 +1,4 @@
-Unsynchronized.set Thy_Output.source;
+Thy_Output.source_default := true;
 use "../../antiquote_setup.ML";
 
 use_thys ["Basics", "Interfaces", "Presentation", "Misc"];
--- a/doc-src/TutorialI/Documents/Documents.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Documents/Documents.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -144,7 +144,7 @@
 definition xor :: "bool \<Rightarrow> bool \<Rightarrow> bool"    (infixl "\<oplus>" 60)
 where "A \<oplus> B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)"
 (*<*)
-local
+setup {* Sign.local_path *}
 (*>*)
 
 text {*
@@ -169,7 +169,7 @@
 
 notation (xsymbols) xor (infixl "\<oplus>\<ignore>" 60)
 (*<*)
-local
+setup {* Sign.local_path *}
 (*>*)
 
 text {*\noindent
--- a/doc-src/TutorialI/Documents/document/Documents.tex	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Documents/document/Documents.tex	Wed Sep 01 17:19:47 2010 +0200
@@ -168,6 +168,19 @@
 \isacommand{definition}\isamarkupfalse%
 \ xor\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}bool\ {\isasymRightarrow}\ bool\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymoplus}{\isachardoublequoteclose}\ {\isadigit{6}}{\isadigit{0}}{\isacharparenright}\isanewline
 \isakeyword{where}\ {\isachardoublequoteopen}A\ {\isasymoplus}\ B\ {\isasymequiv}\ {\isacharparenleft}A\ {\isasymand}\ {\isasymnot}\ B{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymnot}\ A\ {\isasymand}\ B{\isacharparenright}{\isachardoublequoteclose}%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
 \begin{isamarkuptext}%
 \noindent The X-Symbol package within Proof~General provides several
   input methods to enter \isa{{\isasymoplus}} in the text.  If all fails one may
@@ -200,6 +213,19 @@
 \isanewline
 \isacommand{notation}\isamarkupfalse%
 \ {\isacharparenleft}xsymbols{\isacharparenright}\ xor\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymoplus}{\isasymignore}{\isachardoublequoteclose}\ {\isadigit{6}}{\isadigit{0}}{\isacharparenright}%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
 \begin{isamarkuptext}%
 \noindent
 The \commdx{notation} command associates a mixfix
--- a/doc-src/TutorialI/Misc/Itrev.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Misc/Itrev.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -2,7 +2,7 @@
 theory Itrev
 imports Main
 begin
-ML"Unsynchronized.reset unique_names"
+ML"unique_names := false"
 (*>*)
 
 section{*Induction Heuristics*}
@@ -141,6 +141,6 @@
 \index{induction heuristics|)}
 *}
 (*<*)
-ML"Unsynchronized.set unique_names"
+ML"unique_names := true"
 end
 (*>*)
--- a/doc-src/TutorialI/Rules/Basic.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Rules/Basic.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -187,7 +187,7 @@
 
 text{*unification failure trace *}
 
-ML "Unsynchronized.set trace_unify_fail"
+ML "trace_unify_fail := true"
 
 lemma "P(a, f(b, g(e,a), b), a) \<Longrightarrow> P(a, f(b, g(c,a), b), a)"
 txt{*
@@ -212,7 +212,7 @@
 *}
 oops
 
-ML "Unsynchronized.reset trace_unify_fail"
+ML "trace_unify_fail := false"
 
 
 text{*Quantifiers*}
--- a/doc-src/TutorialI/Rules/Primes.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Rules/Primes.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -1,4 +1,3 @@
-(* ID:         $Id$ *)
 (* EXTRACT from HOL/ex/Primes.thy*)
 
 (*Euclid's algorithm 
@@ -10,7 +9,7 @@
 
 
 ML "Pretty.margin_default := 64"
-ML "Thy_Output.indent := 5"  (*that is, Doc/TutorialI/settings.ML*)
+declare [[thy_output_indent = 5]]  (*that is, Doc/TutorialI/settings.ML*)
 
 
 text {*Now in Basic.thy!
--- a/doc-src/TutorialI/Sets/Examples.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Sets/Examples.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -1,7 +1,6 @@
-(* ID:         $Id$ *)
 theory Examples imports Main Binomial begin
 
-ML "Unsynchronized.reset eta_contract"
+ML "eta_contract := false"
 ML "Pretty.margin_default := 64"
 
 text{*membership, intersection *}
--- a/doc-src/TutorialI/Types/Numbers.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Types/Numbers.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -3,7 +3,7 @@
 begin
 
 ML "Pretty.margin_default := 64"
-ML "Thy_Output.indent := 0"  (*we don't want 5 for listing theorems*)
+declare [[thy_output_indent = 0]]  (*we don't want 5 for listing theorems*)
 
 text{*
 
--- a/doc-src/TutorialI/Types/document/Numbers.tex	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Types/document/Numbers.tex	Wed Sep 01 17:19:47 2010 +0200
@@ -26,16 +26,16 @@
 %
 \isatagML
 \isacommand{ML}\isamarkupfalse%
-\ {\isachardoublequoteopen}Pretty{\isachardot}margin{\isacharunderscore}default\ {\isacharcolon}{\isacharequal}\ {\isadigit{6}}{\isadigit{4}}{\isachardoublequoteclose}\isanewline
-\isacommand{ML}\isamarkupfalse%
-\ {\isachardoublequoteopen}Thy{\isacharunderscore}Output{\isachardot}indent\ {\isacharcolon}{\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}%
+\ {\isachardoublequoteopen}Pretty{\isachardot}margin{\isacharunderscore}default\ {\isacharcolon}{\isacharequal}\ {\isadigit{6}}{\isadigit{4}}{\isachardoublequoteclose}%
 \endisatagML
 {\isafoldML}%
 %
 \isadelimML
+\isanewline
 %
 \endisadelimML
-%
+\isacommand{declare}\isamarkupfalse%
+\ {\isacharbrackleft}{\isacharbrackleft}thy{\isacharunderscore}output{\isacharunderscore}indent\ {\isacharequal}\ {\isadigit{0}}{\isacharbrackright}{\isacharbrackright}%
 \begin{isamarkuptext}%
 numeric literals; default simprules; can re-orient%
 \end{isamarkuptext}%
--- a/doc-src/TutorialI/settings.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/settings.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -1,3 +1,1 @@
-(* $Id$ *)
-
-Thy_Output.indent := 5;
+Thy_Output.indent_default := 5;
--- a/doc-src/antiquote_setup.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/antiquote_setup.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -71,8 +71,8 @@
     in
       "\\indexdef{}{" ^ kind' ^ "}{" ^ clean_string txt1 ^ "}" ^
       (txt'
-      |> (if ! Thy_Output.quotes then quote else I)
-      |> (if ! Thy_Output.display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"
+      |> (if Config.get ctxt Thy_Output.quotes then quote else I)
+      |> (if Config.get ctxt Thy_Output.display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"
           else split_lines #> map verbatim #> space_implode "\\isasep\\isanewline%\n"))
     end);
 
@@ -93,18 +93,18 @@
   (Scan.repeat (Attrib.thm -- Scan.lift (Args.parens Args.name)))
   (fn {context = ctxt, ...} =>
     map (apfst (Thy_Output.pretty_thm ctxt))
-    #> (if ! Thy_Output.quotes then map (apfst Pretty.quote) else I)
-    #> (if ! Thy_Output.display
+    #> (if Config.get ctxt Thy_Output.quotes then map (apfst Pretty.quote) else I)
+    #> (if Config.get ctxt Thy_Output.display
         then
           map (fn (p, name) =>
-            Output.output (Pretty.string_of (Pretty.indent (! Thy_Output.indent) p)) ^
-            "\\rulename{" ^ Output.output (Pretty.str_of (Thy_Output.pretty_text name)) ^ "}")
+            Output.output (Pretty.string_of (Pretty.indent (Config.get ctxt Thy_Output.indent) p)) ^
+            "\\rulename{" ^ Output.output (Pretty.str_of (Thy_Output.pretty_text ctxt name)) ^ "}")
           #> space_implode "\\par\\smallskip%\n"
           #> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
         else
           map (fn (p, name) =>
             Output.output (Pretty.str_of p) ^
-            "\\rulename{" ^ Output.output (Pretty.str_of (Thy_Output.pretty_text name)) ^ "}")
+            "\\rulename{" ^ Output.output (Pretty.str_of (Thy_Output.pretty_text ctxt name)) ^ "}")
           #> space_implode "\\par\\smallskip%\n"
           #> enclose "\\isa{" "}"));
 
@@ -112,7 +112,8 @@
 (* theory file *)
 
 val _ = Thy_Output.antiquotation "thy_file" (Scan.lift Args.name)
-  (fn _ => fn name => (Thy_Load.check_thy Path.current name; Thy_Output.output [Pretty.str name]));
+  (fn {context = ctxt, ...} =>
+    fn name => (Thy_Load.check_thy Path.current name; Thy_Output.output ctxt [Pretty.str name]));
 
 
 (* Isabelle/Isar entities (with index) *)
@@ -152,8 +153,9 @@
           idx ^
           (Output.output name
             |> (if markup = "" then I else enclose ("\\" ^ markup ^ "{") "}")
-            |> (if ! Thy_Output.quotes then quote else I)
-            |> (if ! Thy_Output.display then enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
+            |> (if Config.get ctxt Thy_Output.quotes then quote else I)
+            |> (if Config.get ctxt Thy_Output.display
+                then enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
                 else hyper o enclose "\\mbox{\\isa{" "}}"))
         else error ("Bad " ^ kind ^ " " ^ quote name)
       end);
--- a/doc-src/manual.bib	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/manual.bib	Wed Sep 01 17:19:47 2010 +0200
@@ -984,6 +984,14 @@
 
 %O
 
+@TechReport{scala-overview-tech-report,
+  author =       {Martin Odersky and al.},
+  title =        {An Overview of the Scala Programming Language},
+  institution =  {EPFL Lausanne, Switzerland},
+  year =         2004,
+  number =       {IC/2004/64}
+}
+
 @Manual{pvs-language,
   title		= {The {PVS} specification language},
   author	= {S. Owre and N. Shankar and J. M. Rushby},
--- a/doc-src/more_antiquote.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/more_antiquote.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -95,7 +95,7 @@
       |> snd
       |> map_filter (fn (_, (some_thm, proper)) => if proper then some_thm else NONE)
       |> map (holize o no_vars ctxt o AxClass.overload thy);
-  in Thy_Output.output (Thy_Output.maybe_pretty_source (pretty_thm ctxt) src thms) end;
+  in Thy_Output.output ctxt (Thy_Output.maybe_pretty_source pretty_thm ctxt src thms) end;
 
 in
 
@@ -124,12 +124,13 @@
 in
 
 val _ = Thy_Output.antiquotation "code_stmts"
-  (parse_const_terms -- Scan.repeat parse_names -- Scan.lift (Args.parens Args.name))
-  (fn {context = ctxt, ...} => fn ((mk_cs, mk_stmtss), target) =>
+  (parse_const_terms -- Scan.repeat parse_names
+    -- Scan.lift (Args.parens (Args.name -- Scan.option Parse.int)))
+  (fn {context = ctxt, ...} => fn ((mk_cs, mk_stmtss), (target, some_width)) =>
     let val thy = ProofContext.theory_of ctxt in
-      Code_Target.code_of thy
-        target NONE "Example" (mk_cs thy)
+      Code_Target.present_code thy (mk_cs thy)
         (fn naming => maps (fn f => f thy naming) mk_stmtss)
+        target some_width "Example" []
       |> typewriter
     end);
 
--- a/doc-src/rail.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/rail.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -97,8 +97,9 @@
     (idx ^
     (Output.output name
       |> (if markup = "" then I else enclose ("\\" ^ markup ^ "{") "}")
-      |> (if ! Thy_Output.quotes then quote else I)
-      |> (if ! Thy_Output.display then enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
+      |> (if Config.get ctxt Thy_Output.quotes then quote else I)
+      |> (if Config.get ctxt Thy_Output.display
+          then enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
           else hyper o enclose "\\mbox{\\isa{" "}}")), style)
   else ("Bad " ^ kind ^ " " ^ name, false)
   end;
--- a/etc/isar-keywords-ZF.el	Thu Aug 26 13:15:37 2010 +0200
+++ b/etc/isar-keywords-ZF.el	Wed Sep 01 17:19:47 2010 +0200
@@ -1,6 +1,6 @@
 ;;
 ;; Keyword classification tables for Isabelle/Isar.
-;; Generated from Pure + Pure-ProofGeneral + FOL + ZF.
+;; Generated from Pure + FOL + ZF.
 ;; *** DO NOT EDIT *** DO NOT EDIT *** DO NOT EDIT ***
 ;;
 
--- a/etc/isar-keywords.el	Thu Aug 26 13:15:37 2010 +0200
+++ b/etc/isar-keywords.el	Wed Sep 01 17:19:47 2010 +0200
@@ -1,6 +1,6 @@
 ;;
 ;; Keyword classification tables for Isabelle/Isar.
-;; Generated from Pure + Pure-ProofGeneral + HOL + HOLCF + HOL-Boogie + HOL-Nominal + HOL-Statespace.
+;; Generated from Pure + HOL + HOLCF + HOL-Boogie + HOL-Nominal + HOL-Statespace.
 ;; *** DO NOT EDIT *** DO NOT EDIT *** DO NOT EDIT ***
 ;;
 
@@ -245,6 +245,7 @@
     "thus"
     "thy_deps"
     "translations"
+    "try"
     "txt"
     "txt_raw"
     "typ"
@@ -398,6 +399,7 @@
     "thm"
     "thm_deps"
     "thy_deps"
+    "try"
     "typ"
     "unused_thms"
     "value"
--- a/src/CCL/ROOT.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/CCL/ROOT.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -8,6 +8,6 @@
 evaluation to weak head-normal form.
 *)
 
-Unsynchronized.set eta_contract;
+eta_contract := true;
 
 use_thys ["Wfd", "Fix"];
--- a/src/FOLP/IFOLP.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/FOLP/IFOLP.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -63,20 +63,22 @@
 
 syntax "_Proof" :: "[p,o]=>prop"    ("(_ /: _)" [51, 10] 5)
 
-ML {*
-
-(*show_proofs:=true displays the proof terms -- they are ENORMOUS*)
-val show_proofs = Unsynchronized.ref false;
-
-fun proof_tr [p,P] = Const (@{const_name Proof}, dummyT) $ P $ p;
-
-fun proof_tr' [P,p] =
-  if ! show_proofs then Const (@{syntax_const "_Proof"}, dummyT) $ p $ P
-  else P  (*this case discards the proof term*);
+parse_translation {*
+  let fun proof_tr [p, P] = Const (@{const_syntax Proof}, dummyT) $ P $ p
+  in [(@{syntax_const "_Proof"}, proof_tr)] end
 *}
 
-parse_translation {* [(@{syntax_const "_Proof"}, proof_tr)] *}
-print_translation {* [(@{const_syntax Proof}, proof_tr')] *}
+(*show_proofs = true displays the proof terms -- they are ENORMOUS*)
+ML {* val (show_proofs, setup_show_proofs) = Attrib.config_bool "show_proofs" (K false) *}
+setup setup_show_proofs
+
+print_translation (advanced) {*
+  let
+    fun proof_tr' ctxt [P, p] =
+      if Config.get ctxt show_proofs then Const (@{syntax_const "_Proof"}, dummyT) $ p $ P
+      else P
+  in [(@{const_syntax Proof}, proof_tr')] end
+*}
 
 axioms
 
--- a/src/HOL/Auth/Event.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Auth/Event.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -22,14 +22,6 @@
        
 consts 
   bad    :: "agent set"                         -- {* compromised agents *}
-  knows  :: "agent => event list => msg set"
-
-
-text{*The constant "spies" is retained for compatibility's sake*}
-
-abbreviation (input)
-  spies  :: "event list => msg set" where
-  "spies == knows Spy"
 
 text{*Spy has access to his own key for spoof messages, but Server is secure*}
 specification (bad)
@@ -37,9 +29,10 @@
   Server_not_bad [iff]: "Server \<notin> bad"
     by (rule exI [of _ "{Spy}"], simp)
 
-primrec
+primrec knows :: "agent => event list => msg set"
+where
   knows_Nil:   "knows A [] = initState A"
-  knows_Cons:
+| knows_Cons:
     "knows A (ev # evs) =
        (if A = Spy then 
         (case ev of
@@ -62,14 +55,20 @@
   therefore the oops case must use Notes
 *)
 
-consts
-  (*Set of items that might be visible to somebody:
+text{*The constant "spies" is retained for compatibility's sake*}
+
+abbreviation (input)
+  spies  :: "event list => msg set" where
+  "spies == knows Spy"
+
+
+(*Set of items that might be visible to somebody:
     complement of the set of fresh items*)
-  used :: "event list => msg set"
 
-primrec
+primrec used :: "event list => msg set"
+where
   used_Nil:   "used []         = (UN B. parts (initState B))"
-  used_Cons:  "used (ev # evs) =
+| used_Cons:  "used (ev # evs) =
                      (case ev of
                         Says A B X => parts {X} \<union> used evs
                       | Gets A X   => used evs
--- a/src/HOL/Auth/NS_Public_Bad.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Auth/NS_Public_Bad.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -203,7 +203,7 @@
 apply clarify
 apply (frule_tac A' = A in 
        Says_imp_knows_Spy [THEN parts.Inj, THEN unique_NB], auto)
-apply (rename_tac C B' evs3)
+apply (rename_tac evs3 B' C)
 txt{*This is the attack!
 @{subgoals[display,indent=0,margin=65]}
 *}
--- a/src/HOL/Auth/Yahalom.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Auth/Yahalom.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -197,6 +197,7 @@
 apply (erule yahalom.induct,
        drule_tac [7] YM4_analz_knows_Spy, analz_freshK, spy_analz, blast)
 apply (simp only: Says_Server_not_range analz_image_freshK_simps)
+apply safe
 done
 
 lemma analz_insert_freshK:
--- a/src/HOL/Boogie/Tools/boogie_commands.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_commands.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -91,7 +91,7 @@
       | _ => (pair ts, K I))
 
     val discharge = fold (Boogie_VCs.discharge o pair vc_name)
-    fun after_qed [thms] = ProofContext.theory (discharge (vcs ~~ thms))
+    fun after_qed [thms] = ProofContext.background_theory (discharge (vcs ~~ thms))
       | after_qed _ = I
   in
     ProofContext.init_global thy
--- a/src/HOL/Boogie/Tools/boogie_loader.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_loader.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -504,7 +504,7 @@
         in
           Const (@{const_name If}, [@{typ bool}, T, T] ---> T) $ c $ t1 $ t2
         end) ||
-      binexp "implies" (binop @{term "op -->"}) ||
+      binexp "implies" (binop @{term HOL.implies}) ||
       scan_line "distinct" num :|-- scan_count exp >>
         (fn [] => @{term True}
           | ts as (t :: _) => mk_distinct (Term.fastype_of t) ts) ||
--- a/src/HOL/Boogie/Tools/boogie_tactics.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_tactics.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -50,11 +50,11 @@
 
 
 local
-  fun explode_conj (@{term "op &"} $ t $ u) = explode_conj t @ explode_conj u
+  fun explode_conj (@{term HOL.conj} $ t $ u) = explode_conj t @ explode_conj u
     | explode_conj t = [t] 
 
-  fun splt (ts, @{term "op -->"} $ t $ u) = splt (ts @ explode_conj t, u)
-    | splt (ts, @{term "op &"} $ t $ u) = splt (ts, t) @ splt (ts, u)
+  fun splt (ts, @{term HOL.implies} $ t $ u) = splt (ts @ explode_conj t, u)
+    | splt (ts, @{term HOL.conj} $ t $ u) = splt (ts, t) @ splt (ts, u)
     | splt (ts, @{term assert_at} $ _ $ t) = [(ts, t)]
     | splt (_, @{term True}) = []
     | splt tp = [tp]
--- a/src/HOL/Boogie/Tools/boogie_vcs.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_vcs.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -59,12 +59,12 @@
     fun vc_of @{term True} = NONE
       | vc_of (@{term assert_at} $ Free (n, _) $ t) =
           SOME (Assert ((n, t), True))
-      | vc_of (@{term "op -->"} $ @{term True} $ u) = vc_of u
-      | vc_of (@{term "op -->"} $ t $ u) =
+      | vc_of (@{term HOL.implies} $ @{term True} $ u) = vc_of u
+      | vc_of (@{term HOL.implies} $ t $ u) =
           vc_of u |> Option.map (assume t)
-      | vc_of (@{term "op &"} $ (@{term assert_at} $ Free (n, _) $ t) $ u) =
+      | vc_of (@{term HOL.conj} $ (@{term assert_at} $ Free (n, _) $ t) $ u) =
           SOME (vc_of u |> the_default True |> assert (n, t))
-      | vc_of (@{term "op &"} $ t $ u) =
+      | vc_of (@{term HOL.conj} $ t $ u) =
           (case (vc_of t, vc_of u) of
             (NONE, r) => r
           | (l, NONE) => l
@@ -74,9 +74,9 @@
 
 val prop_of_vc =
   let
-    fun mk_conj t u = @{term "op &"} $ t $ u
+    fun mk_conj t u = @{term HOL.conj} $ t $ u
 
-    fun term_of (Assume (t, v)) = @{term "op -->"} $ t $ term_of v
+    fun term_of (Assume (t, v)) = @{term HOL.implies} $ t $ term_of v
       | term_of (Assert ((n, t), v)) =
           mk_conj (@{term assert_at} $ Free (n, @{typ bool}) $ t) (term_of v)
       | term_of (Ignore v) = term_of v
--- a/src/HOL/Code_Evaluation.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Code_Evaluation.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -162,7 +162,7 @@
 subsubsection {* Code generator setup *}
 
 lemmas [code del] = term.recs term.cases term.size
-lemma [code, code del]: "eq_class.eq (t1\<Colon>term) t2 \<longleftrightarrow> eq_class.eq t1 t2" ..
+lemma [code, code del]: "HOL.equal (t1\<Colon>term) t2 \<longleftrightarrow> HOL.equal t1 t2" ..
 
 lemma [code, code del]: "(term_of \<Colon> typerep \<Rightarrow> term) = term_of" ..
 lemma [code, code del]: "(term_of \<Colon> term \<Rightarrow> term) = term_of" ..
--- a/src/HOL/Code_Numeral.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Code_Numeral.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -115,12 +115,12 @@
 lemmas [code del] = code_numeral.recs code_numeral.cases
 
 lemma [code]:
-  "eq_class.eq k l \<longleftrightarrow> eq_class.eq (nat_of k) (nat_of l)"
-  by (cases k, cases l) (simp add: eq)
+  "HOL.equal k l \<longleftrightarrow> HOL.equal (nat_of k) (nat_of l)"
+  by (cases k, cases l) (simp add: equal)
 
 lemma [code nbe]:
-  "eq_class.eq (k::code_numeral) k \<longleftrightarrow> True"
-  by (rule HOL.eq_refl)
+  "HOL.equal (k::code_numeral) k \<longleftrightarrow> True"
+  by (rule equal_refl)
 
 
 subsection {* Code numerals as datatype of ints *}
@@ -301,7 +301,7 @@
   (Haskell "Integer")
   (Scala "BigInt")
 
-code_instance code_numeral :: eq
+code_instance code_numeral :: equal
   (Haskell -)
 
 setup {*
@@ -342,7 +342,7 @@
   (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))")
   (Eval "!(fn n => fn m =>/ if m = 0/ then (0, n) else/ (n div m, n mod m))")
 
-code_const "eq_class.eq \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
   (SML "!((_ : Int.int) = _)")
   (OCaml "Big'_int.eq'_big'_int")
   (Haskell infixl 4 "==")
--- a/src/HOL/Decision_Procs/Approximation.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -3305,7 +3305,7 @@
                              (Const (@{const_name Set.member}, _) $
                               Free (name, _) $ _)) = name
         | variable_of_bound (Const (@{const_name Trueprop}, _) $
-                             (Const (@{const_name "op ="}, _) $
+                             (Const (@{const_name HOL.eq}, _) $
                               Free (name, _) $ _)) = name
         | variable_of_bound t = raise TERM ("variable_of_bound", [t])
 
@@ -3422,7 +3422,7 @@
 
 ML {*
   fun calculated_subterms (@{const Trueprop} $ t) = calculated_subterms t
-    | calculated_subterms (@{const "op -->"} $ _ $ t) = calculated_subterms t
+    | calculated_subterms (@{const HOL.implies} $ _ $ t) = calculated_subterms t
     | calculated_subterms (@{term "op <= :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) = [t1, t2]
     | calculated_subterms (@{term "op < :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) = [t1, t2]
     | calculated_subterms (@{term "op : :: real \<Rightarrow> real set \<Rightarrow> bool"} $ t1 $ 
--- a/src/HOL/Decision_Procs/Cooper.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/Cooper.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -1952,11 +1952,11 @@
         | NONE => error "num_of_term: unsupported dvd")
   | fm_of_term ps vs (@{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) =
       @{code Iff} (fm_of_term ps vs t1, fm_of_term ps vs t2)
-  | fm_of_term ps vs (@{term "op &"} $ t1 $ t2) =
+  | fm_of_term ps vs (@{term HOL.conj} $ t1 $ t2) =
       @{code And} (fm_of_term ps vs t1, fm_of_term ps vs t2)
-  | fm_of_term ps vs (@{term "op |"} $ t1 $ t2) =
+  | fm_of_term ps vs (@{term HOL.disj} $ t1 $ t2) =
       @{code Or} (fm_of_term ps vs t1, fm_of_term ps vs t2)
-  | fm_of_term ps vs (@{term "op -->"} $ t1 $ t2) =
+  | fm_of_term ps vs (@{term HOL.implies} $ t1 $ t2) =
       @{code Imp} (fm_of_term ps vs t1, fm_of_term ps vs t2)
   | fm_of_term ps vs (@{term "Not"} $ t') =
       @{code NOT} (fm_of_term ps vs t')
@@ -2016,7 +2016,7 @@
 
 fun term_bools acc t =
   let
-    val is_op = member (op =) [@{term "op &"}, @{term "op |"}, @{term "op -->"}, @{term "op = :: bool => _"},
+    val is_op = member (op =) [@{term HOL.conj}, @{term HOL.disj}, @{term HOL.implies}, @{term "op = :: bool => _"},
       @{term "op = :: int => _"}, @{term "op < :: int => _"},
       @{term "op <= :: int => _"}, @{term "Not"}, @{term "All :: (int => _) => _"},
       @{term "Ex :: (int => _) => _"}, @{term "True"}, @{term "False"}]
--- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -519,9 +519,9 @@
   val [lt, le] = map (Morphism.term phi) [@{term "op \<sqsubset>"}, @{term "op \<sqsubseteq>"}]
   fun h x t =
    case term_of t of
-     Const(@{const_name "op ="}, _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
+     Const(@{const_name HOL.eq}, _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
                             else Ferrante_Rackoff_Data.Nox
-   | @{term "Not"}$(Const(@{const_name "op ="}, _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
+   | @{term "Not"}$(Const(@{const_name HOL.eq}, _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
                             else Ferrante_Rackoff_Data.Nox
    | b$y$z => if Term.could_unify (b, lt) then
                  if term_of x aconv y then Ferrante_Rackoff_Data.Lt
@@ -771,7 +771,7 @@
       in rth end
     | _ => Thm.reflexive ct)
 
-|  Const(@{const_name "op ="},_)$_$Const(@{const_name Groups.zero},_) =>
+|  Const(@{const_name HOL.eq},_)$_$Const(@{const_name Groups.zero},_) =>
    (case whatis x (Thm.dest_arg1 ct) of
     ("c*x+t",[c,t]) =>
        let
@@ -835,7 +835,7 @@
        val rth = Thm.transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
    in rth end
 
-| Const(@{const_name "op ="},_)$a$b =>
+| Const(@{const_name HOL.eq},_)$a$b =>
    let val (ca,cb) = Thm.dest_binop ct
        val T = ctyp_of_term ca
        val th = instantiate' [SOME T] [SOME ca, SOME cb] eq_iff_diff_eq_0
@@ -844,7 +844,7 @@
               (Semiring_Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
        val rth = Thm.transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
    in rth end
-| @{term "Not"} $(Const(@{const_name "op ="},_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct
+| @{term "Not"} $(Const(@{const_name HOL.eq},_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct
 | _ => Thm.reflexive ct
 end;
 
@@ -852,9 +852,9 @@
  let
   fun h x t =
    case term_of t of
-     Const(@{const_name "op ="}, _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
+     Const(@{const_name HOL.eq}, _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
                             else Ferrante_Rackoff_Data.Nox
-   | @{term "Not"}$(Const(@{const_name "op ="}, _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
+   | @{term "Not"}$(Const(@{const_name HOL.eq}, _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
                             else Ferrante_Rackoff_Data.Nox
    | Const(@{const_name Orderings.less},_)$y$z =>
        if term_of x aconv y then Ferrante_Rackoff_Data.Lt
--- a/src/HOL/Decision_Procs/Ferrack.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/Ferrack.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -1996,9 +1996,9 @@
       @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) 
   | fm_of_term vs (@{term "op \<longleftrightarrow> :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) =
       @{code Iff} (fm_of_term vs t1, fm_of_term vs t2)
-  | fm_of_term vs (@{term "op &"} $ t1 $ t2) = @{code And} (fm_of_term vs t1, fm_of_term vs t2)
-  | fm_of_term vs (@{term "op |"} $ t1 $ t2) = @{code Or} (fm_of_term vs t1, fm_of_term vs t2)
-  | fm_of_term vs (@{term "op -->"} $ t1 $ t2) = @{code Imp} (fm_of_term vs t1, fm_of_term vs t2)
+  | fm_of_term vs (@{term HOL.conj} $ t1 $ t2) = @{code And} (fm_of_term vs t1, fm_of_term vs t2)
+  | fm_of_term vs (@{term HOL.disj} $ t1 $ t2) = @{code Or} (fm_of_term vs t1, fm_of_term vs t2)
+  | fm_of_term vs (@{term HOL.implies} $ t1 $ t2) = @{code Imp} (fm_of_term vs t1, fm_of_term vs t2)
   | fm_of_term vs (@{term "Not"} $ t') = @{code NOT} (fm_of_term vs t')
   | fm_of_term vs (Const (@{const_name Ex}, _) $ Abs (xn, xT, p)) =
       @{code E} (fm_of_term (("", dummyT) :: vs) p)
--- a/src/HOL/Decision_Procs/MIR.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/MIR.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -5837,11 +5837,11 @@
       @{code Dvd} (HOLogic.dest_numeral t1, num_of_term vs t2)
   | fm_of_term vs (@{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) =
       @{code Iff} (fm_of_term vs t1, fm_of_term vs t2)
-  | fm_of_term vs (@{term "op &"} $ t1 $ t2) =
+  | fm_of_term vs (@{term HOL.conj} $ t1 $ t2) =
       @{code And} (fm_of_term vs t1, fm_of_term vs t2)
-  | fm_of_term vs (@{term "op |"} $ t1 $ t2) =
+  | fm_of_term vs (@{term HOL.disj} $ t1 $ t2) =
       @{code Or} (fm_of_term vs t1, fm_of_term vs t2)
-  | fm_of_term vs (@{term "op -->"} $ t1 $ t2) =
+  | fm_of_term vs (@{term HOL.implies} $ t1 $ t2) =
       @{code Imp} (fm_of_term vs t1, fm_of_term vs t2)
   | fm_of_term vs (@{term "Not"} $ t') =
       @{code NOT} (fm_of_term vs t')
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -912,7 +912,7 @@
 
 definition "lt p = (case p of CP (C c) \<Rightarrow> if 0>\<^sub>N c then T else F| _ \<Rightarrow> Lt p)"
 definition "le p = (case p of CP (C c) \<Rightarrow> if 0\<ge>\<^sub>N c then T else F | _ \<Rightarrow> Le p)"
-definition "eq p = (case p of CP (C c) \<Rightarrow> if c = 0\<^sub>N then T else F | _ \<Rightarrow> Eq p)"
+definition eq where "eq p = (case p of CP (C c) \<Rightarrow> if c = 0\<^sub>N then T else F | _ \<Rightarrow> Eq p)"
 definition "neq p = not (eq p)"
 
 lemma lt: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (lt p) = Ifm vs bs (Lt p)"
@@ -2954,13 +2954,13 @@
 fun powt rT = Const(@{const_name "power"}, [rT,@{typ "nat"}] ---> rT);
 val brT = [bT, bT] ---> bT;
 val nott = @{term "Not"};
-val conjt = @{term "op &"};
-val disjt = @{term "op |"};
-val impt = @{term "op -->"};
+val conjt = @{term HOL.conj};
+val disjt = @{term HOL.disj};
+val impt = @{term HOL.implies};
 val ifft = @{term "op = :: bool => _"}
 fun llt rT = Const(@{const_name Orderings.less},rrT rT);
 fun lle rT = Const(@{const_name Orderings.less},rrT rT);
-fun eqt rT = Const(@{const_name "op ="},rrT rT);
+fun eqt rT = Const(@{const_name HOL.eq},rrT rT);
 fun rz rT = Const(@{const_name Groups.zero},rT);
 
 fun dest_nat t = case t of
@@ -3018,10 +3018,10 @@
     Const(@{const_name True},_) => @{code T}
   | Const(@{const_name False},_) => @{code F}
   | Const(@{const_name Not},_)$p => @{code NOT} (fm_of_term m m' p)
-  | Const(@{const_name "op &"},_)$p$q => @{code And} (fm_of_term m m' p, fm_of_term m m' q)
-  | Const(@{const_name "op |"},_)$p$q => @{code Or} (fm_of_term m m' p, fm_of_term m m' q)
-  | Const(@{const_name "op -->"},_)$p$q => @{code Imp} (fm_of_term m m' p, fm_of_term m m' q)
-  | Const(@{const_name "op ="},ty)$p$q => 
+  | Const(@{const_name HOL.conj},_)$p$q => @{code And} (fm_of_term m m' p, fm_of_term m m' q)
+  | Const(@{const_name HOL.disj},_)$p$q => @{code Or} (fm_of_term m m' p, fm_of_term m m' q)
+  | Const(@{const_name HOL.implies},_)$p$q => @{code Imp} (fm_of_term m m' p, fm_of_term m m' q)
+  | Const(@{const_name HOL.eq},ty)$p$q => 
        if domain_type ty = bT then @{code Iff} (fm_of_term m m' p, fm_of_term m m' q)
        else @{code Eq} (@{code Sub} (tm_of_term m m' p, tm_of_term m m' q))
   | Const(@{const_name Orderings.less},_)$p$q => 
--- a/src/HOL/Decision_Procs/commutative_ring_tac.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/commutative_ring_tac.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -65,7 +65,7 @@
 (* reification of the equation *)
 val cr_sort = @{sort "comm_ring_1"};
 
-fun reif_eq thy (eq as Const(@{const_name "op ="}, Type("fun", [T, _])) $ lhs $ rhs) =
+fun reif_eq thy (eq as Const(@{const_name HOL.eq}, Type("fun", [T, _])) $ lhs $ rhs) =
       if Sign.of_sort thy (T, cr_sort) then
         let
           val fs = OldTerm.term_frees eq;
--- a/src/HOL/Decision_Procs/ferrante_rackoff.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/ferrante_rackoff.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -33,12 +33,12 @@
              {isolate_conv = icv, whatis = wi, simpset = simpset}):entry) =
 let
  fun uset (vars as (x::vs)) p = case term_of p of
-   Const(@{const_name "op &"}, _)$ _ $ _ =>
+   Const(@{const_name HOL.conj}, _)$ _ $ _ =>
      let
        val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb
        val (lS,lth) = uset vars l  val (rS, rth) = uset vars r
      in (lS@rS, Drule.binop_cong_rule b lth rth) end
- |  Const(@{const_name "op |"}, _)$ _ $ _ =>
+ |  Const(@{const_name HOL.disj}, _)$ _ $ _ =>
      let
        val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb
        val (lS,lth) = uset vars l  val (rS, rth) = uset vars r
@@ -122,12 +122,12 @@
 
    fun decomp_mpinf fm =
      case term_of fm of
-       Const(@{const_name "op &"},_)$_$_ =>
+       Const(@{const_name HOL.conj},_)$_$_ =>
         let val (p,q) = Thm.dest_binop fm
         in ([p,q], myfwd (minf_conj,pinf_conj, nmi_conj, npi_conj,ld_conj)
                          (Thm.cabs x p) (Thm.cabs x q))
         end
-     | Const(@{const_name "op |"},_)$_$_ =>
+     | Const(@{const_name HOL.disj},_)$_$_ =>
         let val (p,q) = Thm.dest_binop fm
         in ([p,q],myfwd (minf_disj, pinf_disj, nmi_disj, npi_disj,ld_disj)
                          (Thm.cabs x p) (Thm.cabs x q))
@@ -175,15 +175,15 @@
  let
   fun h bounds tm =
    (case term_of tm of
-     Const (@{const_name "op ="}, T) $ _ $ _ =>
+     Const (@{const_name HOL.eq}, T) $ _ $ _ =>
        if domain_type T = HOLogic.boolT then find_args bounds tm
        else Thm.dest_fun2 tm
    | Const (@{const_name Not}, _) $ _ => h bounds (Thm.dest_arg tm)
    | Const (@{const_name All}, _) $ _ => find_body bounds (Thm.dest_arg tm)
    | Const (@{const_name Ex}, _) $ _ => find_body bounds (Thm.dest_arg tm)
-   | Const (@{const_name "op &"}, _) $ _ $ _ => find_args bounds tm
-   | Const (@{const_name "op |"}, _) $ _ $ _ => find_args bounds tm
-   | Const (@{const_name "op -->"}, _) $ _ $ _ => find_args bounds tm
+   | Const (@{const_name HOL.conj}, _) $ _ $ _ => find_args bounds tm
+   | Const (@{const_name HOL.disj}, _) $ _ $ _ => find_args bounds tm
+   | Const (@{const_name HOL.implies}, _) $ _ $ _ => find_args bounds tm
    | Const ("==>", _) $ _ $ _ => find_args bounds tm
    | Const ("==", _) $ _ $ _ => find_args bounds tm
    | Const ("all", _) $ _ => find_body bounds (Thm.dest_arg tm)
--- a/src/HOL/Decision_Procs/langford.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/langford.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -69,28 +69,28 @@
 val all_conjuncts = 
  let fun h acc ct = 
   case term_of ct of
-   @{term "op &"}$_$_ => h (h acc (Thm.dest_arg ct)) (Thm.dest_arg1 ct)
+   @{term HOL.conj}$_$_ => h (h acc (Thm.dest_arg ct)) (Thm.dest_arg1 ct)
   | _ => ct::acc
 in h [] end;
 
 fun conjuncts ct =
  case term_of ct of
-  @{term "op &"}$_$_ => (Thm.dest_arg1 ct)::(conjuncts (Thm.dest_arg ct))
+  @{term HOL.conj}$_$_ => (Thm.dest_arg1 ct)::(conjuncts (Thm.dest_arg ct))
 | _ => [ct];
 
 fun fold1 f = foldr1 (uncurry f);
 
-val list_conj = fold1 (fn c => fn c' => Thm.capply (Thm.capply @{cterm "op &"} c) c') ;
+val list_conj = fold1 (fn c => fn c' => Thm.capply (Thm.capply @{cterm HOL.conj} c) c') ;
 
 fun mk_conj_tab th = 
  let fun h acc th = 
    case prop_of th of
-   @{term "Trueprop"}$(@{term "op &"}$p$q) => 
+   @{term "Trueprop"}$(@{term HOL.conj}$p$q) => 
      h (h acc (th RS conjunct2)) (th RS conjunct1)
   | @{term "Trueprop"}$p => (p,th)::acc
 in fold (Termtab.insert Thm.eq_thm) (h [] th) Termtab.empty end;
 
-fun is_conj (@{term "op &"}$_$_) = true
+fun is_conj (@{term HOL.conj}$_$_) = true
   | is_conj _ = false;
 
 fun prove_conj tab cjs = 
@@ -116,7 +116,7 @@
 fun contains x ct = member (op aconv) (OldTerm.term_frees (term_of ct)) (term_of x);
 
 fun is_eqx x eq = case term_of eq of
-   Const(@{const_name "op ="},_)$l$r => l aconv term_of x orelse r aconv term_of x
+   Const(@{const_name HOL.eq},_)$l$r => l aconv term_of x orelse r aconv term_of x
  | _ => false ;
 
 local 
@@ -176,16 +176,16 @@
  let
   fun h bounds tm =
    (case term_of tm of
-     Const (@{const_name "op ="}, T) $ _ $ _ =>
+     Const (@{const_name HOL.eq}, T) $ _ $ _ =>
        if domain_type T = HOLogic.boolT then find_args bounds tm
        else Thm.dest_fun2 tm
    | Const (@{const_name Not}, _) $ _ => h bounds (Thm.dest_arg tm)
    | Const (@{const_name All}, _) $ _ => find_body bounds (Thm.dest_arg tm)
    | Const ("all", _) $ _ => find_body bounds (Thm.dest_arg tm)
    | Const (@{const_name Ex}, _) $ _ => find_body bounds (Thm.dest_arg tm)
-   | Const (@{const_name "op &"}, _) $ _ $ _ => find_args bounds tm
-   | Const (@{const_name "op |"}, _) $ _ $ _ => find_args bounds tm
-   | Const (@{const_name "op -->"}, _) $ _ $ _ => find_args bounds tm
+   | Const (@{const_name HOL.conj}, _) $ _ $ _ => find_args bounds tm
+   | Const (@{const_name HOL.disj}, _) $ _ $ _ => find_args bounds tm
+   | Const (@{const_name HOL.implies}, _) $ _ $ _ => find_args bounds tm
    | Const ("==>", _) $ _ $ _ => find_args bounds tm
    | Const ("==", _) $ _ $ _ => find_args bounds tm
    | Const (@{const_name Trueprop}, _) $ _ => h bounds (Thm.dest_arg tm)
--- a/src/HOL/HOL.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/HOL.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -30,6 +30,7 @@
   "~~/src/Tools/induct.ML"
   ("~~/src/Tools/induct_tacs.ML")
   ("Tools/recfun_codegen.ML")
+  "Tools/try.ML"
 begin
 
 setup {* Intuitionistic.method_setup @{binding iprover} *}
@@ -57,18 +58,12 @@
   False         :: bool
   Not           :: "bool => bool"                   ("~ _" [40] 40)
 
-setup Sign.root_path
+  conj          :: "[bool, bool] => bool"           (infixr "&" 35)
+  disj          :: "[bool, bool] => bool"           (infixr "|" 30)
+  implies       :: "[bool, bool] => bool"           (infixr "-->" 25)
 
-consts
-  "op &"        :: "[bool, bool] => bool"           (infixr "&" 35)
-  "op |"        :: "[bool, bool] => bool"           (infixr "|" 30)
-  "op -->"      :: "[bool, bool] => bool"           (infixr "-->" 25)
+  eq            :: "['a, 'a] => bool"               (infixl "=" 50)
 
-  "op ="        :: "['a, 'a] => bool"               (infixl "=" 50)
-
-setup Sign.local_path
-
-consts
   The           :: "('a => bool) => 'a"
   All           :: "('a => bool) => bool"           (binder "ALL " 10)
   Ex            :: "('a => bool) => bool"           (binder "EX " 10)
@@ -78,7 +73,7 @@
 subsubsection {* Additional concrete syntax *}
 
 notation (output)
-  "op ="  (infix "=" 50)
+  eq  (infix "=" 50)
 
 abbreviation
   not_equal :: "['a, 'a] => bool"  (infixl "~=" 50) where
@@ -89,15 +84,15 @@
 
 notation (xsymbols)
   Not  ("\<not> _" [40] 40) and
-  "op &"  (infixr "\<and>" 35) and
-  "op |"  (infixr "\<or>" 30) and
-  "op -->"  (infixr "\<longrightarrow>" 25) and
+  conj  (infixr "\<and>" 35) and
+  disj  (infixr "\<or>" 30) and
+  implies  (infixr "\<longrightarrow>" 25) and
   not_equal  (infix "\<noteq>" 50)
 
 notation (HTML output)
   Not  ("\<not> _" [40] 40) and
-  "op &"  (infixr "\<and>" 35) and
-  "op |"  (infixr "\<or>" 30) and
+  conj  (infixr "\<and>" 35) and
+  disj  (infixr "\<or>" 30) and
   not_equal  (infix "\<noteq>" 50)
 
 abbreviation (iff)
@@ -183,8 +178,8 @@
   True_or_False:  "(P=True) | (P=False)"
 
 finalconsts
-  "op ="
-  "op -->"
+  eq
+  implies
   The
 
 definition If :: "bool \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("(if (_)/ then (_)/ else (_))" [0, 0, 10] 10) where
@@ -864,7 +859,7 @@
 
 setup {*
 let
-  fun non_bool_eq (@{const_name "op ="}, Type (_, [T, _])) = T <> @{typ bool}
+  fun non_bool_eq (@{const_name HOL.eq}, Type (_, [T, _])) = T <> @{typ bool}
     | non_bool_eq _ = false;
   val hyp_subst_tac' =
     SUBGOAL (fn (goal, i) =>
@@ -930,7 +925,7 @@
 (
   val thy = @{theory}
   type claset = Classical.claset
-  val equality_name = @{const_name "op ="}
+  val equality_name = @{const_name HOL.eq}
   val not_name = @{const_name Not}
   val notE = @{thm notE}
   val ccontr = @{thm ccontr}
@@ -1578,7 +1573,7 @@
   val atomize_conjL = @{thm atomize_conjL}
   val atomize_disjL = @{thm atomize_disjL}
   val operator_names =
-    [@{const_name "op |"}, @{const_name "op &"}, @{const_name Ex}]
+    [@{const_name HOL.disj}, @{const_name HOL.conj}, @{const_name Ex}]
 );
 *}
 
@@ -1737,8 +1732,8 @@
   "True"    ("true")
   "False"   ("false")
   "Not"     ("Bool.not")
-  "op |"    ("(_ orelse/ _)")
-  "op &"    ("(_ andalso/ _)")
+  HOL.disj    ("(_ orelse/ _)")
+  HOL.conj    ("(_ andalso/ _)")
   "If"      ("(if _/ then _/ else _)")
 
 setup {*
@@ -1746,8 +1741,8 @@
 
 fun eq_codegen thy defs dep thyname b t gr =
     (case strip_comb t of
-       (Const (@{const_name "op ="}, Type (_, [Type ("fun", _), _])), _) => NONE
-     | (Const (@{const_name "op ="}, _), [t, u]) =>
+       (Const (@{const_name HOL.eq}, Type (_, [Type ("fun", _), _])), _) => NONE
+     | (Const (@{const_name HOL.eq}, _), [t, u]) =>
           let
             val (pt, gr') = Codegen.invoke_codegen thy defs dep thyname false t gr;
             val (pu, gr'') = Codegen.invoke_codegen thy defs dep thyname false u gr';
@@ -1756,7 +1751,7 @@
             SOME (Codegen.parens
               (Pretty.block [pt, Codegen.str " =", Pretty.brk 1, pu]), gr''')
           end
-     | (t as Const (@{const_name "op ="}, _), ts) => SOME (Codegen.invoke_codegen
+     | (t as Const (@{const_name HOL.eq}, _), ts) => SOME (Codegen.invoke_codegen
          thy defs dep thyname b (Codegen.eta_expand t ts 2) gr)
      | _ => NONE);
 
@@ -1775,31 +1770,30 @@
 
 subsubsection {* Equality *}
 
-class eq =
-  fixes eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
-  assumes eq_equals: "eq x y \<longleftrightarrow> x = y"
+class equal =
+  fixes equal :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+  assumes equal_eq: "equal x y \<longleftrightarrow> x = y"
 begin
 
-lemma eq [code_unfold, code_inline del]: "eq = (op =)"
-  by (rule ext eq_equals)+
+lemma equal [code_unfold, code_inline del]: "equal = (op =)"
+  by (rule ext equal_eq)+
 
-lemma eq_refl: "eq x x \<longleftrightarrow> True"
-  unfolding eq by rule+
+lemma equal_refl: "equal x x \<longleftrightarrow> True"
+  unfolding equal by rule+
 
-lemma equals_eq: "(op =) \<equiv> eq"
-  by (rule eq_reflection) (rule ext, rule ext, rule sym, rule eq_equals)
-
-declare equals_eq [symmetric, code_post]
+lemma eq_equal: "(op =) \<equiv> equal"
+  by (rule eq_reflection) (rule ext, rule ext, rule sym, rule equal_eq)
 
 end
 
-declare equals_eq [code]
+declare eq_equal [symmetric, code_post]
+declare eq_equal [code]
 
 setup {*
   Code_Preproc.map_pre (fn simpset =>
-    simpset addsimprocs [Simplifier.simproc_global_i @{theory} "eq" [@{term "op ="}]
+    simpset addsimprocs [Simplifier.simproc_global_i @{theory} "equal" [@{term HOL.eq}]
       (fn thy => fn _ => fn Const (_, T) => case strip_type T
-        of (Type _ :: _, _) => SOME @{thm equals_eq}
+        of (Type _ :: _, _) => SOME @{thm eq_equal}
          | _ => NONE)])
 *}
 
@@ -1839,51 +1833,49 @@
     and "(P \<longrightarrow> False) \<longleftrightarrow> \<not> P"
     and "(P \<longrightarrow> True) \<longleftrightarrow> True" by simp_all
 
-instantiation itself :: (type) eq
+instantiation itself :: (type) equal
 begin
 
-definition eq_itself :: "'a itself \<Rightarrow> 'a itself \<Rightarrow> bool" where
-  "eq_itself x y \<longleftrightarrow> x = y"
+definition equal_itself :: "'a itself \<Rightarrow> 'a itself \<Rightarrow> bool" where
+  "equal_itself x y \<longleftrightarrow> x = y"
 
 instance proof
-qed (fact eq_itself_def)
+qed (fact equal_itself_def)
 
 end
 
-lemma eq_itself_code [code]:
-  "eq_class.eq TYPE('a) TYPE('a) \<longleftrightarrow> True"
-  by (simp add: eq)
+lemma equal_itself_code [code]:
+  "equal TYPE('a) TYPE('a) \<longleftrightarrow> True"
+  by (simp add: equal)
 
 text {* Equality *}
 
 declare simp_thms(6) [code nbe]
 
 setup {*
-  Sign.add_const_constraint (@{const_name eq}, SOME @{typ "'a\<Colon>type \<Rightarrow> 'a \<Rightarrow> bool"})
+  Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a\<Colon>type \<Rightarrow> 'a \<Rightarrow> bool"})
 *}
 
-lemma equals_alias_cert: "OFCLASS('a, eq_class) \<equiv> ((op = :: 'a \<Rightarrow> 'a \<Rightarrow> bool) \<equiv> eq)" (is "?ofclass \<equiv> ?eq")
+lemma equal_alias_cert: "OFCLASS('a, equal_class) \<equiv> ((op = :: 'a \<Rightarrow> 'a \<Rightarrow> bool) \<equiv> equal)" (is "?ofclass \<equiv> ?equal")
 proof
   assume "PROP ?ofclass"
-  show "PROP ?eq"
-    by (tactic {* ALLGOALS (rtac (Thm.unconstrainT @{thm equals_eq})) *})
+  show "PROP ?equal"
+    by (tactic {* ALLGOALS (rtac (Thm.unconstrainT @{thm eq_equal})) *})
       (fact `PROP ?ofclass`)
 next
-  assume "PROP ?eq"
+  assume "PROP ?equal"
   show "PROP ?ofclass" proof
-  qed (simp add: `PROP ?eq`)
+  qed (simp add: `PROP ?equal`)
 qed
   
 setup {*
-  Sign.add_const_constraint (@{const_name eq}, SOME @{typ "'a\<Colon>eq \<Rightarrow> 'a \<Rightarrow> bool"})
+  Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a\<Colon>equal \<Rightarrow> 'a \<Rightarrow> bool"})
 *}
 
 setup {*
-  Nbe.add_const_alias @{thm equals_alias_cert}
+  Nbe.add_const_alias @{thm equal_alias_cert}
 *}
 
-hide_const (open) eq
-
 text {* Cases *}
 
 lemma Let_case_cert:
@@ -1904,9 +1896,10 @@
 
 code_abort undefined
 
+
 subsubsection {* Generic code generator target languages *}
 
-text {* type bool *}
+text {* type @{typ bool} *}
 
 code_type bool
   (SML "bool")
@@ -1914,7 +1907,7 @@
   (Haskell "Bool")
   (Scala "Boolean")
 
-code_const True and False and Not and "op &" and "op |" and If
+code_const True and False and Not and HOL.conj and HOL.disj and If
   (SML "true" and "false" and "not"
     and infixl 1 "andalso" and infixl 0 "orelse"
     and "!(if (_)/ then (_)/ else (_))")
@@ -1924,7 +1917,7 @@
   (Haskell "True" and "False" and "not"
     and infixl 3 "&&" and infixl 2 "||"
     and "!(if (_)/ then (_)/ else (_))")
-  (Scala "true" and "false" and "'!/ _"
+  (Scala "true" and "false" and "'! _"
     and infixl 3 "&&" and infixl 1 "||"
     and "!(if ((_))/ (_)/ else (_))")
 
@@ -1939,13 +1932,13 @@
 
 text {* using built-in Haskell equality *}
 
-code_class eq
+code_class equal
   (Haskell "Eq")
 
-code_const "eq_class.eq"
+code_const "HOL.equal"
   (Haskell infixl 4 "==")
 
-code_const "op ="
+code_const HOL.eq
   (Haskell infixl 4 "==")
 
 text {* undefined *}
@@ -2134,4 +2127,6 @@
 
 *}
 
+hide_const (open) eq equal
+
 end
--- a/src/HOL/Imperative_HOL/Heap_Monad.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Imperative_HOL/Heap_Monad.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -477,34 +477,44 @@
 subsubsection {* Scala *}
 
 code_include Scala "Heap"
-{*import collection.mutable.ArraySeq
-import Natural._
-
-def bind[A, B](f: Unit => A, g: A => Unit => B): Unit => B = (_: Unit) => g (f ()) ()
+{*object Heap {
+  def bind[A, B](f: Unit => A, g: A => Unit => B): Unit => B = (_: Unit) => g (f ()) ()
+}
 
 class Ref[A](x: A) {
   var value = x
 }
 
 object Ref {
-  def apply[A](x: A): Ref[A] = new Ref[A](x)
-  def lookup[A](r: Ref[A]): A = r.value
-  def update[A](r: Ref[A], x: A): Unit = { r.value = x }
+  def apply[A](x: A): Ref[A] =
+    new Ref[A](x)
+  def lookup[A](r: Ref[A]): A =
+    r.value
+  def update[A](r: Ref[A], x: A): Unit =
+    { r.value = x }
 }
 
 object Array {
-  def alloc[A](n: Natural)(x: A): ArraySeq[A] = ArraySeq.fill(n.as_Int)(x)
-  def make[A](n: Natural)(f: Natural => A): ArraySeq[A] = ArraySeq.tabulate(n.as_Int)((k: Int) => f(Natural(k)))
-  def len[A](a: ArraySeq[A]): Natural = Natural(a.length)
-  def nth[A](a: ArraySeq[A], n: Natural): A = a(n.as_Int)
-  def upd[A](a: ArraySeq[A], n: Natural, x: A): Unit = a.update(n.as_Int, x)
-  def freeze[A](a: ArraySeq[A]): List[A] = a.toList
-}*}
+  import collection.mutable.ArraySeq
+  def alloc[A](n: Natural)(x: A): ArraySeq[A] =
+    ArraySeq.fill(n.as_Int)(x)
+  def make[A](n: Natural)(f: Natural => A): ArraySeq[A] =
+    ArraySeq.tabulate(n.as_Int)((k: Int) => f(Natural(k)))
+  def len[A](a: ArraySeq[A]): Natural =
+    Natural(a.length)
+  def nth[A](a: ArraySeq[A], n: Natural): A =
+    a(n.as_Int)
+  def upd[A](a: ArraySeq[A], n: Natural, x: A): Unit =
+    a.update(n.as_Int, x)
+  def freeze[A](a: ArraySeq[A]): List[A] =
+    a.toList
+}
+*}
 
-code_reserved Scala bind Ref Array
+code_reserved Scala Heap Ref Array
 
 code_type Heap (Scala "Unit/ =>/ _")
-code_const bind (Scala "bind")
+code_const bind (Scala "Heap.bind")
 code_const return (Scala "('_: Unit)/ =>/ _")
 code_const Heap_Monad.raise' (Scala "!error((_))")
 
--- a/src/HOL/Import/Generate-HOL/GenHOL4Base.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/Generate-HOL/GenHOL4Base.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -17,8 +17,8 @@
   T               > True
   F               > False
   "!"             > All
-  "/\\"           > "op &"
-  "\\/"           > "op |"
+  "/\\"           > HOL.conj
+  "\\/"           > HOL.disj
   "?"             > Ex
   "?!"            > Ex1
   "~"             > Not
--- a/src/HOL/Import/Generate-HOLLight/GenHOLLight.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/Generate-HOLLight/GenHOLLight.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -53,10 +53,10 @@
   F > False
   ONE_ONE > HOL4Setup.ONE_ONE
   ONTO    > Fun.surj
-  "=" > "op ="
-  "==>" > "op -->"
-  "/\\" > "op &"
-  "\\/" > "op |"
+  "=" > HOL.eq
+  "==>" > HOL.implies
+  "/\\" > HOL.conj
+  "\\/" > HOL.disj
   "!" > All
   "?" > Ex
   "?!" > Ex1
--- a/src/HOL/Import/HOL/bool.imp	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/HOL/bool.imp	Wed Sep 01 17:19:47 2010 +0200
@@ -14,7 +14,7 @@
 const_maps
   "~" > "HOL.Not"
   "bool_case" > "Datatype.bool.bool_case"
-  "\\/" > "op |"
+  "\\/" > HOL.disj
   "TYPE_DEFINITION" > "HOL4Setup.TYPE_DEFINITION"
   "T" > "HOL.True"
   "RES_SELECT" > "HOL4Base.bool.RES_SELECT"
@@ -30,7 +30,7 @@
   "ARB" > "HOL4Base.bool.ARB"
   "?!" > "HOL.Ex1"
   "?" > "HOL.Ex"
-  "/\\" > "op &"
+  "/\\" > HOL.conj
   "!" > "HOL.All"
 
 thm_maps
--- a/src/HOL/Import/HOLLight/hollight.imp	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/HOLLight/hollight.imp	Wed Sep 01 17:19:47 2010 +0200
@@ -115,7 +115,7 @@
   "_10303" > "HOLLight.hollight._10303"
   "_10302" > "HOLLight.hollight._10302"
   "_0" > "0" :: "nat"
-  "\\/" > "op |"
+  "\\/" > HOL.disj
   "ZRECSPACE" > "HOLLight.hollight.ZRECSPACE"
   "ZIP" > "HOLLight.hollight.ZIP"
   "ZCONSTR" > "HOLLight.hollight.ZCONSTR"
@@ -233,11 +233,11 @@
   "?" > "HOL.Ex"
   ">=" > "HOLLight.hollight.>="
   ">" > "HOLLight.hollight.>"
-  "==>" > "op -->"
-  "=" > "op ="
+  "==>" > HOL.implies
+  "=" > HOL.eq
   "<=" > "HOLLight.hollight.<="
   "<" > "HOLLight.hollight.<"
-  "/\\" > "op &"
+  "/\\" > HOL.conj
   "-" > "Groups.minus" :: "nat => nat => nat"
   "," > "Product_Type.Pair"
   "+" > "Groups.plus" :: "nat => nat => nat"
--- a/src/HOL/Import/hol4rews.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/hol4rews.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -627,8 +627,8 @@
         thy |> add_hol4_type_mapping "min" "bool" false @{type_name bool}
             |> add_hol4_type_mapping "min" "fun" false "fun"
             |> add_hol4_type_mapping "min" "ind" false @{type_name ind}
-            |> add_hol4_const_mapping "min" "=" false @{const_name "op ="}
-            |> add_hol4_const_mapping "min" "==>" false @{const_name "op -->"}
+            |> add_hol4_const_mapping "min" "=" false @{const_name HOL.eq}
+            |> add_hol4_const_mapping "min" "==>" false @{const_name HOL.implies}
             |> add_hol4_const_mapping "min" "@" false @{const_name "Eps"}
 in
 val hol4_setup =
--- a/src/HOL/Import/import_syntax.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/import_syntax.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -148,11 +148,11 @@
         val _ = TextIO.closeIn is
         val orig_source = Source.of_string inp
         val symb_source = Symbol.source {do_recover = false} orig_source
-        val lexes = Unsynchronized.ref
-          (Scan.make_lexicon (map Symbol.explode ["import_segment","ignore_thms","import","end",">","::","const_maps","const_moves","thm_maps","const_renames","type_maps","def_maps"]),
+        val lexes =
+          (Scan.make_lexicon
+            (map Symbol.explode ["import_segment","ignore_thms","import","end",">","::","const_maps","const_moves","thm_maps","const_renames","type_maps","def_maps"]),
                   Scan.empty_lexicon)
-        val get_lexes = fn () => !lexes
-        val token_source = Token.source {do_recover = NONE} get_lexes Position.start symb_source
+        val token_source = Token.source {do_recover = NONE} (K lexes) Position.start symb_source
         val token_list = filter_out (Token.is_kind Token.Space) (Source.exhaust token_source)
         val import_segmentP = Parse.$$$ "import_segment" |-- import_segment
         val type_mapsP = Parse.$$$ "type_maps" |-- type_maps
--- a/src/HOL/Import/proof_kernel.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/proof_kernel.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -1205,7 +1205,7 @@
 fun non_trivial_term_consts t = fold_aterms
   (fn Const (c, _) =>
       if c = @{const_name Trueprop} orelse c = @{const_name All}
-        orelse c = @{const_name "op -->"} orelse c = @{const_name "op &"} orelse c = @{const_name "op ="}
+        orelse c = @{const_name HOL.implies} orelse c = @{const_name HOL.conj} orelse c = @{const_name HOL.eq}
       then I else insert (op =) c
     | _ => I) t [];
 
@@ -1213,11 +1213,11 @@
     let
         fun add_consts (Const (c, _), cs) =
             (case c of
-                 @{const_name "op ="} => insert (op =) "==" cs
-               | @{const_name "op -->"} => insert (op =) "==>" cs
+                 @{const_name HOL.eq} => insert (op =) "==" cs
+               | @{const_name HOL.implies} => insert (op =) "==>" cs
                | @{const_name All} => cs
                | "all" => cs
-               | @{const_name "op &"} => cs
+               | @{const_name HOL.conj} => cs
                | @{const_name Trueprop} => cs
                | _ => insert (op =) c cs)
           | add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs))
@@ -1476,10 +1476,10 @@
 fun mk_COMB th1 th2 thy =
     let
         val (f,g) = case concl_of th1 of
-                        _ $ (Const(@{const_name "op ="},_) $ f $ g) => (f,g)
+                        _ $ (Const(@{const_name HOL.eq},_) $ f $ g) => (f,g)
                       | _ => raise ERR "mk_COMB" "First theorem not an equality"
         val (x,y) = case concl_of th2 of
-                        _ $ (Const(@{const_name "op ="},_) $ x $ y) => (x,y)
+                        _ $ (Const(@{const_name HOL.eq},_) $ x $ y) => (x,y)
                       | _ => raise ERR "mk_COMB" "Second theorem not an equality"
         val fty = type_of f
         val (fd,fr) = dom_rng fty
@@ -1521,7 +1521,7 @@
         val th1 = norm_hyps th1
         val th2 = norm_hyps th2
         val (l,r) = case concl_of th of
-                        _ $ (Const(@{const_name "op |"},_) $ l $ r) => (l,r)
+                        _ $ (Const(@{const_name HOL.disj},_) $ l $ r) => (l,r)
                       | _ => raise ERR "DISJ_CASES" "Conclusion not a disjunction"
         val th1' = rearrange thy (HOLogic.mk_Trueprop l) th1
         val th2' = rearrange thy (HOLogic.mk_Trueprop r) th2
@@ -1788,7 +1788,7 @@
         val cv = cterm_of thy v
         val th1 = implies_elim_all (beta_eta_thm th)
         val (f,g) = case concl_of th1 of
-                        _ $ (Const(@{const_name "op ="},_) $ f $ g) => (Term.lambda v f,Term.lambda v g)
+                        _ $ (Const(@{const_name HOL.eq},_) $ f $ g) => (Term.lambda v f,Term.lambda v g)
                       | _ => raise ERR "mk_ABS" "Bad conclusion"
         val (fd,fr) = dom_rng (type_of f)
         val abs_thm' = Drule.instantiate' [SOME (ctyp_of thy fd), SOME (ctyp_of thy fr)] [SOME (cterm_of thy f), SOME (cterm_of thy g)] abs_thm
@@ -1860,7 +1860,7 @@
         val _ = if_debug pth hth
         val th1 = implies_elim_all (beta_eta_thm th)
         val a = case concl_of th1 of
-                    _ $ (Const(@{const_name "op -->"},_) $ a $ Const(@{const_name False},_)) => a
+                    _ $ (Const(@{const_name HOL.implies},_) $ a $ Const(@{const_name False},_)) => a
                   | _ => raise ERR "NOT_INTRO" "Conclusion of bad form"
         val ca = cterm_of thy a
         val th2 = Thm.equal_elim (Drule.instantiate' [] [SOME ca] not_intro_thm) th1
--- a/src/HOL/Import/shuffler.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/shuffler.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -60,14 +60,14 @@
 
 fun mk_meta_eq th =
     (case concl_of th of
-         Const(@{const_name Trueprop},_) $ (Const(@{const_name "op ="},_) $ _ $ _) => th RS eq_reflection
+         Const(@{const_name Trueprop},_) $ (Const(@{const_name HOL.eq},_) $ _ $ _) => th RS eq_reflection
        | Const("==",_) $ _ $ _ => th
        | _ => raise THM("Not an equality",0,[th]))
     handle _ => raise THM("Couldn't make meta equality",0,[th])  (* FIXME avoid handle _ *)
 
 fun mk_obj_eq th =
     (case concl_of th of
-         Const(@{const_name Trueprop},_) $ (Const(@{const_name "op ="},_) $ _ $ _) => th
+         Const(@{const_name Trueprop},_) $ (Const(@{const_name HOL.eq},_) $ _ $ _) => th
        | Const("==",_) $ _ $ _ => th RS meta_eq_to_obj_eq
        | _ => raise THM("Not an equality",0,[th]))
     handle _ => raise THM("Couldn't make object equality",0,[th])  (* FIXME avoid handle _ *)
--- a/src/HOL/Int.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Int.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -2222,42 +2222,42 @@
   mult_bin_simps
   arith_extra_simps(4) [where 'a = int]
 
-instantiation int :: eq
+instantiation int :: equal
 begin
 
 definition
-  "eq_class.eq k l \<longleftrightarrow> k - l = (0\<Colon>int)"
-
-instance by default (simp add: eq_int_def)
+  "HOL.equal k l \<longleftrightarrow> k - l = (0\<Colon>int)"
+
+instance by default (simp add: equal_int_def)
 
 end
 
 lemma eq_number_of_int_code [code]:
-  "eq_class.eq (number_of k \<Colon> int) (number_of l) \<longleftrightarrow> eq_class.eq k l"
-  unfolding eq_int_def number_of_is_id ..
+  "HOL.equal (number_of k \<Colon> int) (number_of l) \<longleftrightarrow> HOL.equal k l"
+  unfolding equal_int_def number_of_is_id ..
 
 lemma eq_int_code [code]:
-  "eq_class.eq Int.Pls Int.Pls \<longleftrightarrow> True"
-  "eq_class.eq Int.Pls Int.Min \<longleftrightarrow> False"
-  "eq_class.eq Int.Pls (Int.Bit0 k2) \<longleftrightarrow> eq_class.eq Int.Pls k2"
-  "eq_class.eq Int.Pls (Int.Bit1 k2) \<longleftrightarrow> False"
-  "eq_class.eq Int.Min Int.Pls \<longleftrightarrow> False"
-  "eq_class.eq Int.Min Int.Min \<longleftrightarrow> True"
-  "eq_class.eq Int.Min (Int.Bit0 k2) \<longleftrightarrow> False"
-  "eq_class.eq Int.Min (Int.Bit1 k2) \<longleftrightarrow> eq_class.eq Int.Min k2"
-  "eq_class.eq (Int.Bit0 k1) Int.Pls \<longleftrightarrow> eq_class.eq k1 Int.Pls"
-  "eq_class.eq (Int.Bit1 k1) Int.Pls \<longleftrightarrow> False"
-  "eq_class.eq (Int.Bit0 k1) Int.Min \<longleftrightarrow> False"
-  "eq_class.eq (Int.Bit1 k1) Int.Min \<longleftrightarrow> eq_class.eq k1 Int.Min"
-  "eq_class.eq (Int.Bit0 k1) (Int.Bit0 k2) \<longleftrightarrow> eq_class.eq k1 k2"
-  "eq_class.eq (Int.Bit0 k1) (Int.Bit1 k2) \<longleftrightarrow> False"
-  "eq_class.eq (Int.Bit1 k1) (Int.Bit0 k2) \<longleftrightarrow> False"
-  "eq_class.eq (Int.Bit1 k1) (Int.Bit1 k2) \<longleftrightarrow> eq_class.eq k1 k2"
-  unfolding eq_equals by simp_all
+  "HOL.equal Int.Pls Int.Pls \<longleftrightarrow> True"
+  "HOL.equal Int.Pls Int.Min \<longleftrightarrow> False"
+  "HOL.equal Int.Pls (Int.Bit0 k2) \<longleftrightarrow> HOL.equal Int.Pls k2"
+  "HOL.equal Int.Pls (Int.Bit1 k2) \<longleftrightarrow> False"
+  "HOL.equal Int.Min Int.Pls \<longleftrightarrow> False"
+  "HOL.equal Int.Min Int.Min \<longleftrightarrow> True"
+  "HOL.equal Int.Min (Int.Bit0 k2) \<longleftrightarrow> False"
+  "HOL.equal Int.Min (Int.Bit1 k2) \<longleftrightarrow> HOL.equal Int.Min k2"
+  "HOL.equal (Int.Bit0 k1) Int.Pls \<longleftrightarrow> HOL.equal k1 Int.Pls"
+  "HOL.equal (Int.Bit1 k1) Int.Pls \<longleftrightarrow> False"
+  "HOL.equal (Int.Bit0 k1) Int.Min \<longleftrightarrow> False"
+  "HOL.equal (Int.Bit1 k1) Int.Min \<longleftrightarrow> HOL.equal k1 Int.Min"
+  "HOL.equal (Int.Bit0 k1) (Int.Bit0 k2) \<longleftrightarrow> HOL.equal k1 k2"
+  "HOL.equal (Int.Bit0 k1) (Int.Bit1 k2) \<longleftrightarrow> False"
+  "HOL.equal (Int.Bit1 k1) (Int.Bit0 k2) \<longleftrightarrow> False"
+  "HOL.equal (Int.Bit1 k1) (Int.Bit1 k2) \<longleftrightarrow> HOL.equal k1 k2"
+  unfolding equal_eq by simp_all
 
 lemma eq_int_refl [code nbe]:
-  "eq_class.eq (k::int) k \<longleftrightarrow> True"
-  by (rule HOL.eq_refl)
+  "HOL.equal (k::int) k \<longleftrightarrow> True"
+  by (rule equal_refl)
 
 lemma less_eq_number_of_int_code [code]:
   "(number_of k \<Colon> int) \<le> number_of l \<longleftrightarrow> k \<le> l"
--- a/src/HOL/IsaMakefile	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/IsaMakefile	Wed Sep 01 17:19:47 2010 +0200
@@ -110,6 +110,7 @@
   $(SRC)/Tools/Code/code_eval.ML \
   $(SRC)/Tools/Code/code_haskell.ML \
   $(SRC)/Tools/Code/code_ml.ML \
+  $(SRC)/Tools/Code/code_namespace.ML \
   $(SRC)/Tools/Code/code_preproc.ML \
   $(SRC)/Tools/Code/code_printer.ML \
   $(SRC)/Tools/Code/code_scala.ML \
@@ -213,6 +214,7 @@
   Tools/sat_funcs.ML \
   Tools/sat_solver.ML \
   Tools/split_rule.ML \
+  Tools/try.ML \
   Tools/typedef.ML \
   Transitive_Closure.thy \
   Typedef.thy \
@@ -1321,7 +1323,9 @@
   Predicate_Compile_Examples/ROOT.ML					\
   Predicate_Compile_Examples/Predicate_Compile_Examples.thy		\
   Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy  \
-  Predicate_Compile_Examples/Code_Prolog_Examples.thy
+  Predicate_Compile_Examples/Code_Prolog_Examples.thy 			\
+  Predicate_Compile_Examples/Hotel_Example.thy 				\
+  Predicate_Compile_Examples/Lambda_Example.thy 
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Predicate_Compile_Examples
 
 
--- a/src/HOL/Lazy_Sequence.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Lazy_Sequence.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -39,10 +39,14 @@
   "size xq = (case yield xq of None => 0 | Some (x, xq') => size xq' + 1)"
 by (cases xq) auto
 
-lemma [code]: "eq_class.eq xq yq = (case (yield xq, yield yq) of
-  (None, None) => True | (Some (x, xq'), Some (y, yq')) => (HOL.eq x y) \<and> (eq_class.eq xq yq) | _ => False)"
-apply (cases xq) apply (cases yq) apply (auto simp add: eq_class.eq_equals) 
-apply (cases yq) apply (auto simp add: eq_class.eq_equals) done
+lemma [code]: "HOL.equal xq yq = (case (yield xq, yield yq) of
+  (None, None) => True | (Some (x, xq'), Some (y, yq')) => (HOL.equal x y) \<and> (HOL.equal xq yq) | _ => False)"
+apply (cases xq) apply (cases yq) apply (auto simp add: equal_eq) 
+apply (cases yq) apply (auto simp add: equal_eq) done
+
+lemma [code nbe]:
+  "HOL.equal (x :: 'a lazy_sequence) x \<longleftrightarrow> True"
+  by (fact equal_refl)
 
 lemma seq_case [code]:
   "lazy_sequence_case f g xq = (case (yield xq) of None => f | Some (x, xq') => g x xq')"
--- a/src/HOL/Library/AssocList.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/AssocList.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -701,7 +701,44 @@
   "Mapping.bulkload vs = Mapping (map (\<lambda>n. (n, vs ! n)) [0..<length vs])"
   by (rule mapping_eqI) (simp add: map_of_map_restrict expand_fun_eq)
 
-lemma [code, code del]:
-  "HOL.eq (x :: (_, _) mapping) y \<longleftrightarrow> x = y" by (fact eq_equals) (*FIXME*)
+lemma map_of_eqI: (*FIXME move to Map.thy*)
+  assumes set_eq: "set (map fst xs) = set (map fst ys)"
+  assumes map_eq: "\<forall>k\<in>set (map fst xs). map_of xs k = map_of ys k"
+  shows "map_of xs = map_of ys"
+proof (rule ext)
+  fix k show "map_of xs k = map_of ys k"
+  proof (cases "map_of xs k")
+    case None then have "k \<notin> set (map fst xs)" by (simp add: map_of_eq_None_iff)
+    with set_eq have "k \<notin> set (map fst ys)" by simp
+    then have "map_of ys k = None" by (simp add: map_of_eq_None_iff)
+    with None show ?thesis by simp
+  next
+    case (Some v) then have "k \<in> set (map fst xs)" by (auto simp add: dom_map_of_conv_image_fst [symmetric])
+    with map_eq show ?thesis by auto
+  qed
+qed
+
+lemma map_of_eq_dom: (*FIXME move to Map.thy*)
+  assumes "map_of xs = map_of ys"
+  shows "fst ` set xs = fst ` set ys"
+proof -
+  from assms have "dom (map_of xs) = dom (map_of ys)" by simp
+  then show ?thesis by (simp add: dom_map_of_conv_image_fst)
+qed
+
+lemma equal_Mapping [code]:
+  "HOL.equal (Mapping xs) (Mapping ys) \<longleftrightarrow>
+    (let ks = map fst xs; ls = map fst ys
+    in (\<forall>l\<in>set ls. l \<in> set ks) \<and> (\<forall>k\<in>set ks. k \<in> set ls \<and> map_of xs k = map_of ys k))"
+proof -
+  have aux: "\<And>a b xs. (a, b) \<in> set xs \<Longrightarrow> a \<in> fst ` set xs" by (auto simp add: image_def intro!: bexI)
+  show ?thesis
+    by (auto intro!: map_of_eqI simp add: Let_def equal Mapping_def)
+      (auto dest!: map_of_eq_dom intro: aux)
+qed
+
+lemma [code nbe]:
+  "HOL.equal (x :: ('a, 'b) mapping) x \<longleftrightarrow> True"
+  by (fact equal_refl)
 
 end
--- a/src/HOL/Library/Code_Char.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Code_Char.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -19,7 +19,7 @@
   #> String_Code.add_literal_list_string "Haskell"
 *}
 
-code_instance char :: eq
+code_instance char :: equal
   (Haskell -)
 
 code_reserved SML
@@ -31,7 +31,7 @@
 code_reserved Scala
   char
 
-code_const "eq_class.eq \<Colon> char \<Rightarrow> char \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> char \<Rightarrow> char \<Rightarrow> bool"
   (SML "!((_ : char) = _)")
   (OCaml "!((_ : char) = _)")
   (Haskell infixl 4 "==")
--- a/src/HOL/Library/Code_Integer.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Code_Integer.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -21,7 +21,7 @@
   (Scala "BigInt")
   (Eval "int")
 
-code_instance int :: eq
+code_instance int :: equal
   (Haskell -)
 
 setup {*
@@ -51,14 +51,14 @@
   (SML "IntInf.- ((_), 1)")
   (OCaml "Big'_int.pred'_big'_int")
   (Haskell "!(_/ -/ 1)")
-  (Scala "!(_/ -/ 1)")
+  (Scala "!(_ -/ 1)")
   (Eval "!(_/ -/ 1)")
 
 code_const Int.succ
   (SML "IntInf.+ ((_), 1)")
   (OCaml "Big'_int.succ'_big'_int")
   (Haskell "!(_/ +/ 1)")
-  (Scala "!(_/ +/ 1)")
+  (Scala "!(_ +/ 1)")
   (Eval "!(_/ +/ 1)")
 
 code_const "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int"
@@ -96,7 +96,7 @@
   (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))")
   (Eval "Integer.div'_mod/ (abs _)/ (abs _)")
 
-code_const "eq_class.eq \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
   (SML "!((_ : IntInf.int) = _)")
   (OCaml "Big'_int.eq'_big'_int")
   (Haskell infixl 4 "==")
--- a/src/HOL/Library/Code_Natural.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Code_Natural.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -52,12 +52,11 @@
     | otherwise = (Natural k, Natural l) where (k, l) = quotRem n m;
 };*}
 
+
 code_reserved Haskell Natural
 
-code_include Scala "Natural" {*
-import scala.Math
-
-object Natural {
+code_include Scala "Natural"
+{*object Natural {
 
   def apply(numeral: BigInt): Natural = new Natural(numeral max 0)
   def apply(numeral: Int): Natural = Natural(BigInt(numeral))
@@ -111,7 +110,7 @@
     false Code_Printer.literal_alternative_numeral) ["Haskell", "Scala"]
 *}
 
-code_instance code_numeral :: eq
+code_instance code_numeral :: equal
   (Haskell -)
 
 code_const "op + \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
@@ -130,7 +129,7 @@
   (Haskell "divMod")
   (Scala infixl 8 "/%")
 
-code_const "eq_class.eq \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
   (Haskell infixl 4 "==")
   (Scala infixl 5 "==")
 
--- a/src/HOL/Library/Code_Prolog.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Code_Prolog.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -9,5 +9,10 @@
 uses "~~/src/HOL/Tools/Predicate_Compile/code_prolog.ML"
 begin
 
+section {* Setup for Numerals *}
+
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name number_of}] *}
+setup {* Predicate_Compile_Data.keep_functions [@{const_name number_of}] *}
+
 end
 
--- a/src/HOL/Library/Dlist.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Dlist.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -109,16 +109,20 @@
 
 text {* Equality *}
 
-instantiation dlist :: (eq) eq
+instantiation dlist :: (equal) equal
 begin
 
-definition "HOL.eq dxs dys \<longleftrightarrow> HOL.eq (list_of_dlist dxs) (list_of_dlist dys)"
+definition "HOL.equal dxs dys \<longleftrightarrow> HOL.equal (list_of_dlist dxs) (list_of_dlist dys)"
 
 instance proof
-qed (simp add: eq_dlist_def eq list_of_dlist_inject)
+qed (simp add: equal_dlist_def equal list_of_dlist_inject)
 
 end
 
+lemma [code nbe]:
+  "HOL.equal (dxs :: 'a::equal dlist) dxs \<longleftrightarrow> True"
+  by (fact equal_refl)
+
 
 section {* Induction principle and case distinction *}
 
--- a/src/HOL/Library/Efficient_Nat.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Efficient_Nat.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -55,12 +55,12 @@
   by (simp add: prod_fun_def split_def pdivmod_def nat_div_distrib nat_mod_distrib divmod_nat_div_mod)
 
 lemma eq_nat_code [code]:
-  "eq_class.eq n m \<longleftrightarrow> eq_class.eq (of_nat n \<Colon> int) (of_nat m)"
-  by (simp add: eq)
+  "HOL.equal n m \<longleftrightarrow> HOL.equal (of_nat n \<Colon> int) (of_nat m)"
+  by (simp add: equal)
 
 lemma eq_nat_refl [code nbe]:
-  "eq_class.eq (n::nat) n \<longleftrightarrow> True"
-  by (rule HOL.eq_refl)
+  "HOL.equal (n::nat) n \<longleftrightarrow> True"
+  by (rule equal_refl)
 
 lemma less_eq_nat_code [code]:
   "n \<le> m \<longleftrightarrow> (of_nat n \<Colon> int) \<le> of_nat m"
@@ -242,8 +242,8 @@
   and @{typ int}.
 *}
 
-code_include Haskell "Nat" {*
-newtype Nat = Nat Integer deriving (Eq, Show, Read);
+code_include Haskell "Nat"
+{*newtype Nat = Nat Integer deriving (Eq, Show, Read);
 
 instance Num Nat where {
   fromInteger k = Nat (if k >= 0 then k else 0);
@@ -280,10 +280,8 @@
 
 code_reserved Haskell Nat
 
-code_include Scala "Nat" {*
-import scala.Math
-
-object Nat {
+code_include Scala "Nat"
+{*object Nat {
 
   def apply(numeral: BigInt): Nat = new Nat(numeral max 0)
   def apply(numeral: Int): Nat = Nat(BigInt(numeral))
@@ -332,7 +330,7 @@
   (Haskell "Nat.Nat")
   (Scala "Nat")
 
-code_instance nat :: eq
+code_instance nat :: equal
   (Haskell -)
 
 text {*
@@ -442,7 +440,7 @@
   (Scala infixl 8 "/%")
   (Eval "Integer.div'_mod")
 
-code_const "eq_class.eq \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"
   (SML "!((_ : IntInf.int) = _)")
   (OCaml "Big'_int.eq'_big'_int")
   (Haskell infixl 4 "==")
--- a/src/HOL/Library/Enum.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Enum.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -35,17 +35,21 @@
 
 subsection {* Equality and order on functions *}
 
-instantiation "fun" :: (enum, eq) eq
+instantiation "fun" :: (enum, equal) equal
 begin
 
 definition
-  "eq_class.eq f g \<longleftrightarrow> (\<forall>x \<in> set enum. f x = g x)"
+  "HOL.equal f g \<longleftrightarrow> (\<forall>x \<in> set enum. f x = g x)"
 
 instance proof
-qed (simp_all add: eq_fun_def enum_all expand_fun_eq)
+qed (simp_all add: equal_fun_def enum_all expand_fun_eq)
 
 end
 
+lemma [code nbe]:
+  "HOL.equal (f :: _ \<Rightarrow> _) f \<longleftrightarrow> True"
+  by (fact equal_refl)
+
 lemma order_fun [code]:
   fixes f g :: "'a\<Colon>enum \<Rightarrow> 'b\<Colon>order"
   shows "f \<le> g \<longleftrightarrow> list_all (\<lambda>x. f x \<le> g x) enum"
@@ -169,7 +173,7 @@
 
 end
 
-lemma enum_fun_code [code]: "enum = (let enum_a = (enum \<Colon> 'a\<Colon>{enum, eq} list)
+lemma enum_fun_code [code]: "enum = (let enum_a = (enum \<Colon> 'a\<Colon>{enum, equal} list)
   in map (\<lambda>ys. the o map_of (zip enum_a ys)) (n_lists (length enum_a) enum))"
   by (simp add: enum_fun_def Let_def)
 
--- a/src/HOL/Library/Fset.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Fset.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -227,17 +227,21 @@
   "A < B \<longleftrightarrow> A \<le> B \<and> \<not> B \<le> (A :: 'a fset)"
   by (fact less_le_not_le)
 
-instantiation fset :: (type) eq
+instantiation fset :: (type) equal
 begin
 
 definition
-  "eq_fset A B \<longleftrightarrow> A \<le> B \<and> B \<le> (A :: 'a fset)"
+  "HOL.equal A B \<longleftrightarrow> A \<le> B \<and> B \<le> (A :: 'a fset)"
 
 instance proof
-qed (simp add: eq_fset_def set_eq [symmetric])
+qed (simp add: equal_fset_def set_eq [symmetric])
 
 end
 
+lemma [code nbe]:
+  "HOL.equal (A :: 'a fset) A \<longleftrightarrow> True"
+  by (fact equal_refl)
+
 
 subsection {* Functorial operations *}
 
--- a/src/HOL/Library/List_Prefix.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/List_Prefix.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -81,9 +81,9 @@
   by (auto simp add: prefix_def)
 
 lemma less_eq_list_code [code]:
-  "([]\<Colon>'a\<Colon>{eq, ord} list) \<le> xs \<longleftrightarrow> True"
-  "(x\<Colon>'a\<Colon>{eq, ord}) # xs \<le> [] \<longleftrightarrow> False"
-  "(x\<Colon>'a\<Colon>{eq, ord}) # xs \<le> y # ys \<longleftrightarrow> x = y \<and> xs \<le> ys"
+  "([]\<Colon>'a\<Colon>{equal, ord} list) \<le> xs \<longleftrightarrow> True"
+  "(x\<Colon>'a\<Colon>{equal, ord}) # xs \<le> [] \<longleftrightarrow> False"
+  "(x\<Colon>'a\<Colon>{equal, ord}) # xs \<le> y # ys \<longleftrightarrow> x = y \<and> xs \<le> ys"
   by simp_all
 
 lemma same_prefix_prefix [simp]: "(xs @ ys \<le> xs @ zs) = (ys \<le> zs)"
--- a/src/HOL/Library/List_lexord.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/List_lexord.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -103,15 +103,15 @@
 end
 
 lemma less_list_code [code]:
-  "xs < ([]\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> False"
-  "[] < (x\<Colon>'a\<Colon>{eq, order}) # xs \<longleftrightarrow> True"
-  "(x\<Colon>'a\<Colon>{eq, order}) # xs < y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs < ys"
+  "xs < ([]\<Colon>'a\<Colon>{equal, order} list) \<longleftrightarrow> False"
+  "[] < (x\<Colon>'a\<Colon>{equal, order}) # xs \<longleftrightarrow> True"
+  "(x\<Colon>'a\<Colon>{equal, order}) # xs < y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs < ys"
   by simp_all
 
 lemma less_eq_list_code [code]:
-  "x # xs \<le> ([]\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> False"
-  "[] \<le> (xs\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> True"
-  "(x\<Colon>'a\<Colon>{eq, order}) # xs \<le> y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs \<le> ys"
+  "x # xs \<le> ([]\<Colon>'a\<Colon>{equal, order} list) \<longleftrightarrow> False"
+  "[] \<le> (xs\<Colon>'a\<Colon>{equal, order} list) \<longleftrightarrow> True"
+  "(x\<Colon>'a\<Colon>{equal, order}) # xs \<le> y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs \<le> ys"
   by simp_all
 
 end
--- a/src/HOL/Library/Mapping.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Mapping.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -280,14 +280,14 @@
 
 code_datatype empty update
 
-instantiation mapping :: (type, type) eq
+instantiation mapping :: (type, type) equal
 begin
 
 definition [code del]:
-  "HOL.eq m n \<longleftrightarrow> lookup m = lookup n"
+  "HOL.equal m n \<longleftrightarrow> lookup m = lookup n"
 
 instance proof
-qed (simp add: eq_mapping_def)
+qed (simp add: equal_mapping_def)
 
 end
 
--- a/src/HOL/Library/Multiset.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Multiset.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -938,17 +938,21 @@
   qed
 qed
 
-instantiation multiset :: (eq) eq
+instantiation multiset :: (equal) equal
 begin
 
 definition
-  "HOL.eq A B \<longleftrightarrow> (A::'a multiset) \<le> B \<and> B \<le> A"
+  "HOL.equal A B \<longleftrightarrow> (A::'a multiset) \<le> B \<and> B \<le> A"
 
 instance proof
-qed (simp add: eq_multiset_def eq_iff)
+qed (simp add: equal_multiset_def eq_iff)
 
 end
 
+lemma [code nbe]:
+  "HOL.equal (A :: 'a::equal multiset) A \<longleftrightarrow> True"
+  by (fact equal_refl)
+
 definition (in term_syntax)
   bagify :: "('a\<Colon>typerep \<times> nat) list \<times> (unit \<Rightarrow> Code_Evaluation.term)
     \<Rightarrow> 'a multiset \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
--- a/src/HOL/Library/Nested_Environment.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Nested_Environment.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -521,22 +521,21 @@
 text {* Environments and code generation *}
 
 lemma [code, code del]:
-  fixes e1 e2 :: "('b\<Colon>eq, 'a\<Colon>eq, 'c\<Colon>eq) env"
-  shows "eq_class.eq e1 e2 \<longleftrightarrow> eq_class.eq e1 e2" ..
+  "(HOL.equal :: (_, _, _) env \<Rightarrow> _) = HOL.equal" ..
 
-lemma eq_env_code [code]:
-  fixes x y :: "'a\<Colon>eq"
-    and f g :: "'c\<Colon>{eq, finite} \<Rightarrow> ('b\<Colon>eq, 'a, 'c) env option"
-  shows "eq_class.eq (Env x f) (Env y g) \<longleftrightarrow>
-  eq_class.eq x y \<and> (\<forall>z\<in>UNIV. case f z
+lemma equal_env_code [code]:
+  fixes x y :: "'a\<Colon>equal"
+    and f g :: "'c\<Colon>{equal, finite} \<Rightarrow> ('b\<Colon>equal, 'a, 'c) env option"
+  shows "HOL.equal (Env x f) (Env y g) \<longleftrightarrow>
+  HOL.equal x y \<and> (\<forall>z\<in>UNIV. case f z
    of None \<Rightarrow> (case g z
         of None \<Rightarrow> True | Some _ \<Rightarrow> False)
     | Some a \<Rightarrow> (case g z
-        of None \<Rightarrow> False | Some b \<Rightarrow> eq_class.eq a b))" (is ?env)
-    and "eq_class.eq (Val a) (Val b) \<longleftrightarrow> eq_class.eq a b"
-    and "eq_class.eq (Val a) (Env y g) \<longleftrightarrow> False"
-    and "eq_class.eq (Env x f) (Val b) \<longleftrightarrow> False"
-proof (unfold eq)
+        of None \<Rightarrow> False | Some b \<Rightarrow> HOL.equal a b))" (is ?env)
+    and "HOL.equal (Val a) (Val b) \<longleftrightarrow> HOL.equal a b"
+    and "HOL.equal (Val a) (Env y g) \<longleftrightarrow> False"
+    and "HOL.equal (Env x f) (Val b) \<longleftrightarrow> False"
+proof (unfold equal)
   have "f = g \<longleftrightarrow> (\<forall>z. case f z
    of None \<Rightarrow> (case g z
         of None \<Rightarrow> True | Some _ \<Rightarrow> False)
@@ -562,6 +561,10 @@
           of None \<Rightarrow> False | Some b \<Rightarrow> a = b))" by simp
 qed simp_all
 
+lemma [code nbe]:
+  "HOL.equal (x :: (_, _, _) env) x \<longleftrightarrow> True"
+  by (fact equal_refl)
+
 lemma [code, code del]:
   "(Code_Evaluation.term_of :: ('a::{term_of, type}, 'b::{term_of, type}, 'c::{term_of, type}) env \<Rightarrow> term) = Code_Evaluation.term_of" ..
 
--- a/src/HOL/Library/OptionalSugar.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/OptionalSugar.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -32,7 +32,7 @@
 (* deprecated, use thm with style instead, will be removed *)
 (* aligning equations *)
 notation (tab output)
-  "op ="  ("(_) \<^raw:}\putisatab\isa{\ >=\<^raw:}\putisatab\isa{> (_)" [50,49] 50) and
+  "HOL.eq"  ("(_) \<^raw:}\putisatab\isa{\ >=\<^raw:}\putisatab\isa{> (_)" [50,49] 50) and
   "=="  ("(_) \<^raw:}\putisatab\isa{\ >\<equiv>\<^raw:}\putisatab\isa{> (_)")
 
 (* Let *)
--- a/src/HOL/Library/Polynomial.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Polynomial.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -1505,23 +1505,27 @@
 
 declare pCons_0_0 [code_post]
 
-instantiation poly :: ("{zero,eq}") eq
+instantiation poly :: ("{zero, equal}") equal
 begin
 
 definition
-  "eq_class.eq (p::'a poly) q \<longleftrightarrow> p = q"
+  "HOL.equal (p::'a poly) q \<longleftrightarrow> p = q"
 
-instance
-  by default (rule eq_poly_def)
+instance proof
+qed (rule equal_poly_def)
 
 end
 
 lemma eq_poly_code [code]:
-  "eq_class.eq (0::_ poly) (0::_ poly) \<longleftrightarrow> True"
-  "eq_class.eq (0::_ poly) (pCons b q) \<longleftrightarrow> eq_class.eq 0 b \<and> eq_class.eq 0 q"
-  "eq_class.eq (pCons a p) (0::_ poly) \<longleftrightarrow> eq_class.eq a 0 \<and> eq_class.eq p 0"
-  "eq_class.eq (pCons a p) (pCons b q) \<longleftrightarrow> eq_class.eq a b \<and> eq_class.eq p q"
-unfolding eq by simp_all
+  "HOL.equal (0::_ poly) (0::_ poly) \<longleftrightarrow> True"
+  "HOL.equal (0::_ poly) (pCons b q) \<longleftrightarrow> HOL.equal 0 b \<and> HOL.equal 0 q"
+  "HOL.equal (pCons a p) (0::_ poly) \<longleftrightarrow> HOL.equal a 0 \<and> HOL.equal p 0"
+  "HOL.equal (pCons a p) (pCons b q) \<longleftrightarrow> HOL.equal a b \<and> HOL.equal p q"
+  by (simp_all add: equal)
+
+lemma [code nbe]:
+  "HOL.equal (p :: _ poly) p \<longleftrightarrow> True"
+  by (fact equal_refl)
 
 lemmas coeff_code [code] =
   coeff_0 coeff_pCons_0 coeff_pCons_Suc
--- a/src/HOL/Library/Product_ord.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Product_ord.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -22,8 +22,8 @@
 end
 
 lemma [code]:
-  "(x1\<Colon>'a\<Colon>{ord, eq}, y1) \<le> (x2, y2) \<longleftrightarrow> x1 < x2 \<or> x1 \<le> x2 \<and> y1 \<le> y2"
-  "(x1\<Colon>'a\<Colon>{ord, eq}, y1) < (x2, y2) \<longleftrightarrow> x1 < x2 \<or> x1 \<le> x2 \<and> y1 < y2"
+  "(x1\<Colon>'a\<Colon>{ord, equal}, y1) \<le> (x2, y2) \<longleftrightarrow> x1 < x2 \<or> x1 \<le> x2 \<and> y1 \<le> y2"
+  "(x1\<Colon>'a\<Colon>{ord, equal}, y1) < (x2, y2) \<longleftrightarrow> x1 < x2 \<or> x1 \<le> x2 \<and> y1 < y2"
   unfolding prod_le_def prod_less_def by simp_all
 
 instance prod :: (preorder, preorder) preorder proof
--- a/src/HOL/Library/RBT.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/RBT.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -222,12 +222,14 @@
   "Mapping.bulkload vs = Mapping (bulkload (List.map (\<lambda>n. (n, vs ! n)) [0..<length vs]))"
   by (rule mapping_eqI) (simp add: map_of_map_restrict expand_fun_eq)
 
-lemma [code, code del]:
-  "HOL.eq (x :: (_, _) mapping) y \<longleftrightarrow> x = y" by (fact eq_equals) (*FIXME*)
+lemma equal_Mapping [code]:
+  "HOL.equal (Mapping t1) (Mapping t2) \<longleftrightarrow> entries t1 = entries t2"
+  by (simp add: equal Mapping_def entries_lookup)
 
-lemma eq_Mapping [code]:
-  "HOL.eq (Mapping t1) (Mapping t2) \<longleftrightarrow> entries t1 = entries t2"
-  by (simp add: eq Mapping_def entries_lookup)
+lemma [code nbe]:
+  "HOL.equal (x :: (_, _) mapping) x \<longleftrightarrow> True"
+  by (fact equal_refl)
+
 
 hide_const (open) impl_of lookup empty insert delete
   entries keys bulkload map_entry map fold
--- a/src/HOL/Library/Sum_Of_Squares.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Sum_Of_Squares.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -28,6 +28,7 @@
   without calling an external prover.
 *}
 
+setup Sum_Of_Squares.setup
 setup SOS_Wrapper.setup
 
 text {* Tests *}
--- a/src/HOL/Library/Sum_Of_Squares/sos_wrapper.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Sum_Of_Squares/sos_wrapper.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -8,8 +8,8 @@
 sig
   datatype prover_result = Success | Failure | Error
   val setup: theory -> theory
-  val destdir: string Unsynchronized.ref
-  val prover_name: string Unsynchronized.ref
+  val dest_dir: string Config.T
+  val prover_name: string Config.T
 end
 
 structure SOS_Wrapper: SOS_WRAPPER =
@@ -22,14 +22,9 @@
   | str_of_result Error = "Error"
 
 
-(*** output control ***)
-
-fun debug s = if ! Sum_Of_Squares.debugging then writeln s else ()
-
-
 (*** calling provers ***)
 
-val destdir = Unsynchronized.ref ""
+val (dest_dir, setup_dest_dir) = Attrib.config_string "sos_dest_dir" (K "")
 
 fun filename dir name =
   let
@@ -43,12 +38,12 @@
     else error ("No such directory: " ^ dir)
   end
 
-fun run_solver name cmd find_failure input =
+fun run_solver ctxt name cmd find_failure input =
   let
     val _ = warning ("Calling solver: " ^ name)
 
     (* create input file *)
-    val dir = ! destdir
+    val dir = Config.get ctxt dest_dir
     val input_file = filename dir "sos_in"
     val _ = File.write input_file input
 
@@ -71,7 +66,10 @@
         (File.rm input_file; if File.exists output_file then File.rm output_file else ())
       else ()
 
-    val _ = debug ("Solver output:\n" ^ output)
+    val _ =
+      if Config.get ctxt Sum_Of_Squares.trace
+      then writeln ("Solver output:\n" ^ output)
+      else ()
 
     val _ = warning (str_of_result res ^ ": " ^ res_msg)
   in
@@ -120,13 +118,13 @@
   | prover "csdp" = csdp
   | prover name = error ("Unknown prover: " ^ name)
 
-val prover_name = Unsynchronized.ref "remote_csdp"
+val (prover_name, setup_prover_name) = Attrib.config_string "sos_prover_name" (K "remote_csdp")
 
-fun call_solver opt_name =
+fun call_solver ctxt opt_name =
   let
-    val name = the_default (! prover_name) opt_name
+    val name = the_default (Config.get ctxt prover_name) opt_name
     val (cmd, find_failure) = prover name
-  in run_solver name (Path.explode cmd) find_failure end
+  in run_solver ctxt name (Path.explode cmd) find_failure end
 
 
 (* certificate output *)
@@ -143,9 +141,13 @@
 fun sos_solver print method = SIMPLE_METHOD' o Sum_Of_Squares.sos_tac print method
 
 val setup =
+  setup_dest_dir #>
+  setup_prover_name #>
   Method.setup @{binding sos}
     (Scan.lift (Scan.option Parse.xname)
-      >> (sos_solver print_cert o Sum_Of_Squares.Prover o call_solver))
+      >> (fn opt_name => fn ctxt =>
+        sos_solver print_cert
+          (Sum_Of_Squares.Prover (call_solver ctxt opt_name)) ctxt))
     "prove universal problems over the reals using sums of squares" #>
   Method.setup @{binding sos_cert}
     (Scan.lift Parse.string
--- a/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -9,7 +9,8 @@
 sig
   datatype proof_method = Certificate of RealArith.pss_tree | Prover of string -> string
   val sos_tac: (RealArith.pss_tree -> unit) -> proof_method -> Proof.context -> int -> tactic
-  val debugging: bool Unsynchronized.ref
+  val trace: bool Config.T
+  val setup: theory -> theory
   exception Failure of string;
 end
 
@@ -49,7 +50,8 @@
 val pow2 = rat_pow rat_2;
 val pow10 = rat_pow rat_10;
 
-val debugging = Unsynchronized.ref false;
+val (trace, setup_trace) = Attrib.config_bool "sos_trace" (K false);
+val setup = setup_trace;
 
 exception Sanity;
 
@@ -1059,7 +1061,7 @@
 (* Positiv- and Nullstellensatz. Flag "linf" forces a linear representation. *)
 
 
-fun real_positivnullstellensatz_general prover linf d eqs leqs pol =
+fun real_positivnullstellensatz_general ctxt prover linf d eqs leqs pol =
 let
  val vars = fold_rev (union (op aconvc) o poly_variables)
    (pol :: eqs @ map fst leqs) []
@@ -1129,7 +1131,7 @@
   fun find_rounding d =
    let
     val _ =
-      if !debugging
+      if Config.get ctxt trace
       then writeln ("Trying rounding with limit "^Rat.string_of_rat d ^ "\n")
       else ()
     val vec = nice_vector d raw_vec
@@ -1245,7 +1247,7 @@
            let val e = multidegree pol
                val k = if e = 0 then 0 else d div e
                val eq' = map fst eq
-           in tryfind (fn i => (d,i,real_positivnullstellensatz_general prover false d eq' leq
+           in tryfind (fn i => (d,i,real_positivnullstellensatz_general ctxt prover false d eq' leq
                                  (poly_neg(poly_pow pol i))))
                    (0 upto k)
            end
@@ -1356,7 +1358,7 @@
 
 val known_sos_constants =
   [@{term "op ==>"}, @{term "Trueprop"},
-   @{term "op -->"}, @{term "op &"}, @{term "op |"},
+   @{term HOL.implies}, @{term HOL.conj}, @{term HOL.disj},
    @{term "Not"}, @{term "op = :: bool => _"},
    @{term "All :: (real => _) => _"}, @{term "Ex :: (real => _) => _"},
    @{term "op = :: real => _"}, @{term "op < :: real => _"},
--- a/src/HOL/Library/positivstellensatz.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/positivstellensatz.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -164,21 +164,6 @@
   thm list * thm list * thm list -> thm * pss_tree
 type cert_conv = cterm -> thm * pss_tree
 
-val my_eqs = Unsynchronized.ref ([] : thm list);
-val my_les = Unsynchronized.ref ([] : thm list);
-val my_lts = Unsynchronized.ref ([] : thm list);
-val my_proof = Unsynchronized.ref (Axiom_eq 0);
-val my_context = Unsynchronized.ref @{context};
-
-val my_mk_numeric = Unsynchronized.ref ((K @{cterm True}) :Rat.rat -> cterm);
-val my_numeric_eq_conv = Unsynchronized.ref no_conv;
-val my_numeric_ge_conv = Unsynchronized.ref no_conv;
-val my_numeric_gt_conv = Unsynchronized.ref no_conv;
-val my_poly_conv = Unsynchronized.ref no_conv;
-val my_poly_neg_conv = Unsynchronized.ref no_conv;
-val my_poly_add_conv = Unsynchronized.ref no_conv;
-val my_poly_mul_conv = Unsynchronized.ref no_conv;
-
 
     (* Some useful derived rules *)
 fun deduct_antisym_rule tha thb = 
@@ -362,11 +347,6 @@
        poly_conv,poly_neg_conv,poly_add_conv,poly_mul_conv,
        absconv1,absconv2,prover) = 
 let
- val _ = my_context := ctxt 
- val _ = (my_mk_numeric := mk_numeric ; my_numeric_eq_conv := numeric_eq_conv ; 
-          my_numeric_ge_conv := numeric_ge_conv; my_numeric_gt_conv := numeric_gt_conv ;
-          my_poly_conv := poly_conv; my_poly_neg_conv := poly_neg_conv; 
-          my_poly_add_conv := poly_add_conv; my_poly_mul_conv := poly_mul_conv)
  val pre_ss = HOL_basic_ss addsimps simp_thms@ ex_simps@ all_simps@[@{thm not_all},@{thm not_ex},ex_disj_distrib, all_conj_distrib, @{thm if_bool_eq_disj}]
  val prenex_ss = HOL_basic_ss addsimps prenex_simps
  val skolemize_ss = HOL_basic_ss addsimps [choice_iff]
@@ -408,7 +388,6 @@
 
   fun hol_of_positivstellensatz(eqs,les,lts) proof =
    let 
-    val _ = (my_eqs := eqs ; my_les := les ; my_lts := lts ; my_proof := proof)
     fun translate prf = case prf of
         Axiom_eq n => nth eqs n
       | Axiom_le n => nth les n
@@ -439,8 +418,8 @@
   val is_req = is_binop @{cterm "op =:: real => _"}
   val is_ge = is_binop @{cterm "op <=:: real => _"}
   val is_gt = is_binop @{cterm "op <:: real => _"}
-  val is_conj = is_binop @{cterm "op &"}
-  val is_disj = is_binop @{cterm "op |"}
+  val is_conj = is_binop @{cterm HOL.conj}
+  val is_disj = is_binop @{cterm HOL.disj}
   fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2})
   fun disj_cases th th1 th2 = 
    let val (p,q) = Thm.dest_binop (concl th)
@@ -484,7 +463,7 @@
     val th_p = poly_conv(Thm.dest_arg(Thm.dest_arg1(Thm.rhs_of th)))
     val th_x = Drule.arg_cong_rule @{cterm "uminus :: real => _"} th_p
     val th_n = fconv_rule (arg_conv poly_neg_conv) th_x
-    val th' = Drule.binop_cong_rule @{cterm "op |"} 
+    val th' = Drule.binop_cong_rule @{cterm HOL.disj} 
      (Drule.arg_cong_rule (Thm.capply @{cterm "op <::real=>_"} @{cterm "0::real"}) th_p)
      (Drule.arg_cong_rule (Thm.capply @{cterm "op <::real=>_"} @{cterm "0::real"}) th_n)
     in Thm.transitive th th' 
@@ -542,12 +521,12 @@
   let 
    val nnf_norm_conv' = 
      nnf_conv then_conv 
-     literals_conv [@{term "op &"}, @{term "op |"}] [] 
+     literals_conv [@{term HOL.conj}, @{term HOL.disj}] [] 
      (Conv.cache_conv 
        (first_conv [real_lt_conv, real_le_conv, 
                     real_eq_conv, real_not_lt_conv, 
                     real_not_le_conv, real_not_eq_conv, all_conv]))
-  fun absremover ct = (literals_conv [@{term "op &"}, @{term "op |"}] [] 
+  fun absremover ct = (literals_conv [@{term HOL.conj}, @{term HOL.disj}] [] 
                   (try_conv (absconv1 then_conv binop_conv (arg_conv poly_conv))) then_conv 
         try_conv (absconv2 then_conv nnf_norm_conv' then_conv binop_conv absremover)) ct
   val nct = Thm.capply @{cterm Trueprop} (Thm.capply @{cterm "Not"} ct)
--- a/src/HOL/Library/reflection.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/reflection.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -82,7 +82,7 @@
 fun rearrange congs =
   let
     fun P (_, th) =
-      let val @{term "Trueprop"}$(Const (@{const_name "op ="},_) $l$_) = concl_of th
+      let val @{term "Trueprop"}$(Const (@{const_name HOL.eq},_) $l$_) = concl_of th
       in can dest_Var l end
     val (yes,no) = List.partition P congs
   in no @ yes end
--- a/src/HOL/List.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/List.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -4721,8 +4721,8 @@
   by (simp add: null_def)
 
 lemma equal_Nil_null [code_unfold]:
-  "eq_class.eq xs [] \<longleftrightarrow> null xs"
-  by (simp add: eq eq_Nil_null)
+  "HOL.equal xs [] \<longleftrightarrow> null xs"
+  by (simp add: equal eq_Nil_null)
 
 definition maps :: "('a \<Rightarrow> 'b list) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
   [code_post]: "maps f xs = concat (map f xs)"
@@ -4821,10 +4821,10 @@
   (Haskell "[]")
   (Scala "!Nil")
 
-code_instance list :: eq
+code_instance list :: equal
   (Haskell -)
 
-code_const "eq_class.eq \<Colon> 'a\<Colon>eq list \<Rightarrow> 'a list \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
   (Haskell infixl 4 "==")
 
 code_reserved SML
--- a/src/HOL/Matrix/Compute_Oracle/compute.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Matrix/Compute_Oracle/compute.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -371,7 +371,7 @@
 fun merge_shyps shyps1 shyps2 = Sorttab.keys (add_shyps shyps2 (add_shyps shyps1 Sorttab.empty))
 
 val (_, export_oracle) = Context.>>> (Context.map_theory_result
-  (Thm.add_oracle (Binding.name "compute", fn (thy, hyps, shyps, prop) =>
+  (Thm.add_oracle (@{binding compute}, fn (thy, hyps, shyps, prop) =>
     let
         val shyptab = add_shyps shyps Sorttab.empty
         fun delete s shyptab = Sorttab.delete s shyptab handle Sorttab.UNDEF _ => shyptab
--- a/src/HOL/Mirabelle/Mirabelle_Test.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Mirabelle/Mirabelle_Test.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -12,6 +12,7 @@
   "Tools/mirabelle_quickcheck.ML"
   "Tools/mirabelle_refute.ML"
   "Tools/mirabelle_sledgehammer.ML"
+  "Tools/mirabelle_sledgehammer_filter.ML"
 begin
 
 text {*
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -290,10 +290,12 @@
     | NONE => get_prover (default_atp_name ()))
   end
 
+type locality = Sledgehammer_Fact_Filter.locality
+
 local
 
 datatype sh_result =
-  SH_OK of int * int * (string * bool) list |
+  SH_OK of int * int * (string * locality) list |
   SH_FAIL of int * int |
   SH_ERROR
 
@@ -355,15 +357,16 @@
     case result of
       SH_OK (time_isa, time_atp, names) =>
         let
-          fun get_thms (name, chained) =
-            ((name, chained), thms_of_name (Proof.context_of st) name)
+          fun get_thms (name, Sledgehammer_Fact_Filter.Chained) = NONE
+            | get_thms (name, loc) =
+              SOME ((name, loc), thms_of_name (Proof.context_of st) name)
         in
           change_data id inc_sh_success;
           change_data id (inc_sh_lemmas (length names));
           change_data id (inc_sh_max_lems (length names));
           change_data id (inc_sh_time_isa time_isa);
           change_data id (inc_sh_time_atp time_atp);
-          named_thms := SOME (map get_thms names);
+          named_thms := SOME (map_filter get_thms names);
           log (sh_tag id ^ "succeeded (" ^ string_of_int time_isa ^ "+" ^
             string_of_int time_atp ^ ") [" ^ prover_name ^ "]:\n" ^ msg)
         end
@@ -445,7 +448,7 @@
     then () else
     let
       val named_thms =
-        Unsynchronized.ref (NONE : ((string * bool) * thm list) list option)
+        Unsynchronized.ref (NONE : ((string * locality) * thm list) list option)
       val minimize = AList.defined (op =) args minimizeK
       val metis_ft = AList.defined (op =) args metis_ftK
   
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -0,0 +1,169 @@
+(*  Title:      HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML
+    Author:     Jasmin Blanchette, TU Munich
+*)
+
+structure Mirabelle_Sledgehammer_Filter : MIRABELLE_ACTION =
+struct
+
+val relevance_filter_args =
+  [("worse_irrel_freq", Sledgehammer_Fact_Filter.worse_irrel_freq),
+   ("higher_order_irrel_weight",
+    Sledgehammer_Fact_Filter.higher_order_irrel_weight),
+   ("abs_rel_weight", Sledgehammer_Fact_Filter.abs_rel_weight),
+   ("abs_irrel_weight", Sledgehammer_Fact_Filter.abs_irrel_weight),
+   ("skolem_irrel_weight", Sledgehammer_Fact_Filter.skolem_irrel_weight),
+   ("intro_bonus", Sledgehammer_Fact_Filter.intro_bonus),
+   ("elim_bonus", Sledgehammer_Fact_Filter.elim_bonus),
+   ("simp_bonus", Sledgehammer_Fact_Filter.simp_bonus),
+   ("local_bonus", Sledgehammer_Fact_Filter.local_bonus),
+   ("chained_bonus", Sledgehammer_Fact_Filter.chained_bonus),
+   ("max_imperfect", Sledgehammer_Fact_Filter.max_imperfect),
+   ("max_imperfect_exp", Sledgehammer_Fact_Filter.max_imperfect_exp),
+   ("threshold_divisor", Sledgehammer_Fact_Filter.threshold_divisor),
+   ("ridiculous_threshold", Sledgehammer_Fact_Filter.ridiculous_threshold)]
+
+structure Prooftab =
+  Table(type key = int * int val ord = prod_ord int_ord int_ord);
+
+val proof_table = Unsynchronized.ref Prooftab.empty
+
+val num_successes = Unsynchronized.ref ([] : (int * int) list)
+val num_failures = Unsynchronized.ref ([] : (int * int) list)
+val num_found_proofs = Unsynchronized.ref ([] : (int * int) list)
+val num_lost_proofs = Unsynchronized.ref ([] : (int * int) list)
+val num_found_facts = Unsynchronized.ref ([] : (int * int) list)
+val num_lost_facts = Unsynchronized.ref ([] : (int * int) list)
+
+fun get id c = the_default 0 (AList.lookup (op =) (!c) id)
+fun add id c n =
+  c := (case AList.lookup (op =) (!c) id of
+          SOME m => AList.update (op =) (id, m + n) (!c)
+        | NONE => (id, n) :: !c)
+
+fun init proof_file _ thy =
+  let
+    fun do_line line =
+      case line |> space_explode ":" of
+        [line_num, col_num, proof] =>
+        SOME (pairself (the o Int.fromString) (line_num, col_num),
+              proof |> space_explode " " |> filter_out (curry (op =) ""))
+       | _ => NONE
+    val proofs = File.read (Path.explode proof_file)
+    val proof_tab =
+      proofs |> space_explode "\n"
+             |> map_filter do_line
+             |> AList.coalesce (op =)
+             |> Prooftab.make
+  in proof_table := proof_tab; thy end
+
+fun percentage a b = if b = 0 then "N/A" else string_of_int (a * 100 div b)
+fun percentage_alt a b = percentage a (a + b)
+
+fun done id ({log, ...} : Mirabelle.done_args) =
+  if get id num_successes + get id num_failures > 0 then
+    (log "";
+     log ("Number of overall successes: " ^
+          string_of_int (get id num_successes));
+     log ("Number of overall failures: " ^ string_of_int (get id num_failures));
+     log ("Overall success rate: " ^
+          percentage_alt (get id num_successes) (get id num_failures) ^ "%");
+     log ("Number of found proofs: " ^ string_of_int (get id num_found_proofs));
+     log ("Number of lost proofs: " ^ string_of_int (get id num_lost_proofs));
+     log ("Proof found rate: " ^
+          percentage_alt (get id num_found_proofs) (get id num_lost_proofs) ^
+          "%");
+     log ("Number of found facts: " ^ string_of_int (get id num_found_facts));
+     log ("Number of lost facts: " ^ string_of_int (get id num_lost_facts));
+     log ("Fact found rate: " ^
+          percentage_alt (get id num_found_facts) (get id num_lost_facts) ^
+          "%"))
+  else
+    ()
+
+val default_max_relevant = 300
+
+fun with_index (i, s) = s ^ "@" ^ string_of_int i
+
+fun action args id ({pre, pos, log, ...} : Mirabelle.run_args) =
+  case (Position.line_of pos, Position.column_of pos) of
+    (SOME line_num, SOME col_num) =>
+    (case Prooftab.lookup (!proof_table) (line_num, col_num) of
+       SOME proofs =>
+       let
+         val {context = ctxt, facts, goal} = Proof.goal pre
+         val thy = ProofContext.theory_of ctxt
+         val args =
+           args
+           |> filter (fn (key, value) =>
+                         case AList.lookup (op =) relevance_filter_args key of
+                           SOME rf => (rf := the (Real.fromString value); false)
+                         | NONE => true)
+
+         val {relevance_thresholds, full_types, max_relevant, theory_relevant,
+              ...} = Sledgehammer_Isar.default_params thy args
+         val subgoal = 1
+         val (_, hyp_ts, concl_t) = Sledgehammer_Util.strip_subgoal goal subgoal
+         val facts =
+           Sledgehammer_Fact_Filter.relevant_facts ctxt full_types
+               relevance_thresholds
+               (the_default default_max_relevant max_relevant)
+               (the_default false theory_relevant)
+               {add = [], del = [], only = false} facts hyp_ts concl_t
+           |> map (fst o fst)
+         val (found_facts, lost_facts) =
+           List.concat proofs |> sort_distinct string_ord
+           |> map (fn fact => (find_index (curry (op =) fact) facts, fact))
+           |> List.partition (curry (op <=) 0 o fst)
+           |>> sort (prod_ord int_ord string_ord) ||> map snd
+         val found_proofs = filter (forall (member (op =) facts)) proofs
+         val n = length found_proofs
+         val _ =
+           if n = 0 then
+             (add id num_failures 1; log "Failure")
+           else
+             (add id num_successes 1;
+              add id num_found_proofs n;
+              log ("Success (" ^ string_of_int n ^ " of " ^
+                   string_of_int (length proofs) ^ " proofs)"))
+         val _ = add id num_lost_proofs (length proofs - n)
+         val _ = add id num_found_facts (length found_facts)
+         val _ = add id num_lost_facts (length lost_facts)
+         val _ =
+           if null found_facts then
+             ()
+           else
+             let
+               val found_weight =
+                 Real.fromInt (fold (fn (n, _) =>
+                                        Integer.add (n * n)) found_facts 0)
+                   / Real.fromInt (length found_facts)
+                 |> Math.sqrt |> Real.ceil
+             in
+               log ("Found facts (among " ^ string_of_int (length facts) ^
+                    ", weight " ^ string_of_int found_weight ^ "): " ^
+                    commas (map with_index found_facts))
+             end
+         val _ = if null lost_facts then
+                   ()
+                 else
+                   log ("Lost facts (among " ^ string_of_int (length facts) ^
+                        "): " ^ commas lost_facts)
+       in () end
+     | NONE => log "No known proof")
+  | _ => ()
+
+val proof_fileK = "proof_file"
+
+fun invoke args =
+  let
+    val (pf_args, other_args) =
+      args |> List.partition (curry (op =) proof_fileK o fst)
+    val proof_file = case pf_args of
+                       [] => error "No \"proof_file\" specified"
+                     | (_, s) :: _ => s
+  in Mirabelle.register (init proof_file, action other_args, done) end
+
+end;
+
+(* Workaround to keep the "mirabelle.pl" script happy *)
+structure Mirabelle_Sledgehammer_filter = Mirabelle_Sledgehammer_Filter;
--- a/src/HOL/Mirabelle/lib/scripts/mirabelle.pl	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Mirabelle/lib/scripts/mirabelle.pl	Wed Sep 01 17:19:47 2010 +0200
@@ -51,7 +51,11 @@
 }
 my $tools = "";
 if ($#action_files >= 0) {
-  $tools = "uses " . join(" ", @action_files);
+  # uniquify
+  my $s = join ("\n", @action_files);
+  my @action_files = split(/\n/, $s . "\n" . $s);
+  %action_files = sort(@action_files);
+  $tools = "uses " . join(" ", sort(keys(%action_files)));
 }
 
 open(SETUP_FILE, ">$setup_file") || die "Could not create file '$setup_file'";
@@ -71,7 +75,7 @@
 
 END
 
-foreach (split(/:/, $actions)) {
+foreach (reverse(split(/:/, $actions))) {
   if (m/([^[]*)(?:\[(.*)\])?/) {
     my ($name, $settings_str) = ($1, $2 || "");
     $name =~ s/^([a-z])/\U$1/;
--- a/src/HOL/Multivariate_Analysis/Gauge_Measure.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Gauge_Measure.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -311,7 +311,7 @@
   shows "(\<Union> f) has_gmeasure (setsum m f)" using assms
 proof induct case (insert x s)
   have *:"(x \<inter> \<Union>s) = \<Union>{x \<inter> y| y. y\<in>s}"by auto
-  show ?case unfolding Union_insert ring_class.setsum.insert[OF insert(1-2)] 
+  show ?case unfolding Union_insert setsum.insert [OF insert(1-2)] 
     apply(rule has_gmeasure_negligible_union) unfolding *
     apply(rule insert) defer apply(rule insert) apply(rule insert) defer
     apply(rule insert) prefer 4 apply(rule negligible_unions)
--- a/src/HOL/Mutabelle/Mutabelle.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Mutabelle/Mutabelle.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -4,7 +4,7 @@
 begin
 
 ML {*
-val comms = [@{const_name "op ="}, @{const_name "op |"}, @{const_name "op &"}];
+val comms = [@{const_name HOL.eq}, @{const_name HOL.disj}, @{const_name HOL.conj}];
 
 val forbidden =
  [(@{const_name Power.power}, "'a"),
--- a/src/HOL/Mutabelle/mutabelle_extra.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Mutabelle/mutabelle_extra.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -187,7 +187,7 @@
 val nitpick_mtd = ("nitpick", invoke_nitpick)
 *)
 
-val comms = [@{const_name "op ="}, @{const_name "op |"}, @{const_name "op &"}]
+val comms = [@{const_name HOL.eq}, @{const_name HOL.disj}, @{const_name HOL.conj}]
 
 val forbidden =
  [(* (@{const_name "power"}, "'a"), *)
@@ -202,7 +202,7 @@
   (@{const_name "top_fun_inst.top_fun"}, "'a"),
   (@{const_name "Pure.term"}, "'a"),
   (@{const_name "top_class.top"}, "'a"),
-  (@{const_name "eq_class.eq"}, "'a"),
+  (@{const_name "HOL.equal"}, "'a"),
   (@{const_name "Quotient.Quot_True"}, "'a")(*,
   (@{const_name "uminus"}, "'a"),
   (@{const_name "Nat.size"}, "'a"),
@@ -237,7 +237,7 @@
  @{const_name "top_fun_inst.top_fun"},
  @{const_name "Pure.term"},
  @{const_name "top_class.top"},
- @{const_name "eq_class.eq"},
+ @{const_name "HOL.equal"},
  @{const_name "Quotient.Quot_True"}]
 
 fun is_forbidden_mutant t =
--- a/src/HOL/Nominal/nominal_datatype.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Nominal/nominal_datatype.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -183,7 +183,7 @@
   end;
 
 fun mk_not_sym ths = maps (fn th => case prop_of th of
-    _ $ (Const (@{const_name Not}, _) $ (Const (@{const_name "op ="}, _) $ _ $ _)) => [th, th RS not_sym]
+    _ $ (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ _)) => [th, th RS not_sym]
   | _ => [th]) ths;
 
 fun fresh_const T U = Const ("Nominal.fresh", T --> U --> HOLogic.boolT);
@@ -1200,7 +1200,7 @@
           (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs);
     val tnames = Datatype_Prop.make_tnames recTs;
     val zs = Name.variant_list tnames (replicate (length descr'') "z");
-    val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+    val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
       (map (fn ((((i, _), T), tname), z) =>
         make_pred fsT i T $ Free (z, fsT) $ Free (tname, T))
         (descr'' ~~ recTs ~~ tnames ~~ zs)));
@@ -1215,7 +1215,7 @@
         map (make_ind_prem fsT' (fn T => fn t => fn u => HOLogic.Not $
           HOLogic.mk_mem (t, the (AList.lookup op = fresh_fs T) $ u)) i T)
             (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs);
-    val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+    val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
       (map (fn ((((i, _), T), tname), z) =>
         make_pred fsT' i T $ Free (z, fsT') $ Free (tname, T))
         (descr'' ~~ recTs ~~ tnames ~~ zs)));
@@ -1225,7 +1225,7 @@
       (Datatype_Prop.indexify_names (replicate (length dt_atomTs) "pi") ~~
        map mk_permT dt_atomTs) @ [("z", fsT')];
     val aux_ind_Ts = rev (map snd aux_ind_vars);
-    val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+    val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
       (map (fn (((i, _), T), tname) =>
         HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $
           fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1))
--- a/src/HOL/Nominal/nominal_inductive.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Nominal/nominal_inductive.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -71,7 +71,7 @@
   | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs
   | add_binders thy i _ bs = bs;
 
-fun split_conj f names (Const (@{const_name "op &"}, _) $ p $ q) _ = (case head_of p of
+fun split_conj f names (Const (@{const_name HOL.conj}, _) $ p $ q) _ = (case head_of p of
       Const (name, _) =>
         if member (op =) names name then SOME (f p q) else NONE
     | _ => NONE)
@@ -89,7 +89,7 @@
 (* where "id" protects the subformula from simplification            *)
 (*********************************************************************)
 
-fun inst_conj_all names ps pis (Const (@{const_name "op &"}, _) $ p $ q) _ =
+fun inst_conj_all names ps pis (Const (@{const_name HOL.conj}, _) $ p $ q) _ =
       (case head_of p of
          Const (name, _) =>
            if member (op =) names name then SOME (HOLogic.mk_conj (p,
--- a/src/HOL/Nominal/nominal_inductive2.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Nominal/nominal_inductive2.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -76,7 +76,7 @@
   | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs
   | add_binders thy i _ bs = bs;
 
-fun split_conj f names (Const (@{const_name "op &"}, _) $ p $ q) _ = (case head_of p of
+fun split_conj f names (Const (@{const_name HOL.conj}, _) $ p $ q) _ = (case head_of p of
       Const (name, _) =>
         if member (op =) names name then SOME (f p q) else NONE
     | _ => NONE)
@@ -94,7 +94,7 @@
 (* where "id" protects the subformula from simplification            *)
 (*********************************************************************)
 
-fun inst_conj_all names ps pis (Const (@{const_name "op &"}, _) $ p $ q) _ =
+fun inst_conj_all names ps pis (Const (@{const_name HOL.conj}, _) $ p $ q) _ =
       (case head_of p of
          Const (name, _) =>
            if member (op =) names name then SOME (HOLogic.mk_conj (p,
--- a/src/HOL/Nominal/nominal_thmdecls.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Nominal/nominal_thmdecls.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -18,8 +18,6 @@
   val eqvt_force_del: attribute
   val setup: theory -> theory
   val get_eqvt_thms: Proof.context -> thm list
-
-  val NOMINAL_EQVT_DEBUG : bool Unsynchronized.ref
 end;
 
 structure NominalThmDecls: NOMINAL_THMDECLS =
@@ -44,29 +42,31 @@
 (* equality-lemma can be derived. *)
 exception EQVT_FORM of string
 
-val NOMINAL_EQVT_DEBUG = Unsynchronized.ref false
+val (nominal_eqvt_debug, setup_nominal_eqvt_debug) =
+  Attrib.config_bool "nominal_eqvt_debug" (K false);
 
-fun tactic (msg, tac) =
-  if !NOMINAL_EQVT_DEBUG
+fun tactic ctxt (msg, tac) =
+  if Config.get ctxt nominal_eqvt_debug
   then tac THEN' (K (print_tac ("after " ^ msg)))
   else tac
 
 fun prove_eqvt_tac ctxt orig_thm pi pi' =
-let
-  val mypi = Thm.cterm_of ctxt pi
-  val T = fastype_of pi'
-  val mypifree = Thm.cterm_of ctxt (Const (@{const_name "rev"}, T --> T) $ pi')
-  val perm_pi_simp = PureThy.get_thms ctxt "perm_pi_simp"
-in
-  EVERY1 [tactic ("iffI applied", rtac @{thm iffI}),
-          tactic ("remove pi with perm_boolE", dtac @{thm perm_boolE}),
-          tactic ("solve with orig_thm", etac orig_thm),
-          tactic ("applies orig_thm instantiated with rev pi",
-                     dtac (Drule.cterm_instantiate [(mypi,mypifree)] orig_thm)),
-          tactic ("getting rid of the pi on the right", rtac @{thm perm_boolI}),
-          tactic ("getting rid of all remaining perms",
-                     full_simp_tac (HOL_basic_ss addsimps perm_pi_simp))]
-end;
+  let
+    val thy = ProofContext.theory_of ctxt
+    val mypi = Thm.cterm_of thy pi
+    val T = fastype_of pi'
+    val mypifree = Thm.cterm_of thy (Const (@{const_name "rev"}, T --> T) $ pi')
+    val perm_pi_simp = PureThy.get_thms thy "perm_pi_simp"
+  in
+    EVERY1 [tactic ctxt ("iffI applied", rtac @{thm iffI}),
+            tactic ctxt ("remove pi with perm_boolE", dtac @{thm perm_boolE}),
+            tactic ctxt ("solve with orig_thm", etac orig_thm),
+            tactic ctxt ("applies orig_thm instantiated with rev pi",
+                       dtac (Drule.cterm_instantiate [(mypi,mypifree)] orig_thm)),
+            tactic ctxt ("getting rid of the pi on the right", rtac @{thm perm_boolI}),
+            tactic ctxt ("getting rid of all remaining perms",
+                       full_simp_tac (HOL_basic_ss addsimps perm_pi_simp))]
+  end;
 
 fun get_derived_thm ctxt hyp concl orig_thm pi typi =
   let
@@ -78,7 +78,7 @@
     val _ = writeln (Syntax.string_of_term ctxt' goal_term);
   in
     Goal.prove ctxt' [] [] goal_term
-      (fn _ => prove_eqvt_tac thy orig_thm pi' pi'') |>
+      (fn _ => prove_eqvt_tac ctxt' orig_thm pi' pi'') |>
     singleton (ProofContext.export ctxt' ctxt)
   end
 
@@ -147,7 +147,7 @@
              else raise EQVT_FORM "Type Implication"
           end
        (* case: eqvt-lemma is of the equational form *)
-      | (Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, _) $
+      | (Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $
             (Const (@{const_name "perm"},typrm) $ Var (pi,typi) $ lhs) $ rhs)) =>
            (if (apply_pi lhs (pi,typi)) = rhs
                then [orig_thm]
@@ -170,11 +170,12 @@
 val get_eqvt_thms = Context.Proof #> Data.get;
 
 val setup =
-    Attrib.setup @{binding eqvt} (Attrib.add_del eqvt_add eqvt_del) 
-     "equivariance theorem declaration" 
- #> Attrib.setup @{binding eqvt_force} (Attrib.add_del eqvt_force_add eqvt_force_del)
-     "equivariance theorem declaration (without checking the form of the lemma)" 
- #> PureThy.add_thms_dynamic (Binding.name "eqvts", Data.get) 
+  setup_nominal_eqvt_debug #>
+  Attrib.setup @{binding eqvt} (Attrib.add_del eqvt_add eqvt_del) 
+   "equivariance theorem declaration" #>
+  Attrib.setup @{binding eqvt_force} (Attrib.add_del eqvt_force_add eqvt_force_del)
+    "equivariance theorem declaration (without checking the form of the lemma)" #>
+  PureThy.add_thms_dynamic (@{binding eqvts}, Data.get);
 
 
 end;
--- a/src/HOL/Option.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Option.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -99,8 +99,8 @@
   by (simp add: is_none_def)
 
 lemma [code_unfold]:
-  "eq_class.eq x None \<longleftrightarrow> is_none x"
-  by (simp add: eq is_none_none)
+  "HOL.equal x None \<longleftrightarrow> is_none x"
+  by (simp add: equal is_none_none)
 
 hide_const (open) is_none
 
@@ -116,10 +116,10 @@
   (Haskell "Nothing" and "Just")
   (Scala "!None" and "Some")
 
-code_instance option :: eq
+code_instance option :: equal
   (Haskell -)
 
-code_const "eq_class.eq \<Colon> 'a\<Colon>eq option \<Rightarrow> 'a option \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> 'a option \<Rightarrow> 'a option \<Rightarrow> bool"
   (Haskell infixl 4 "==")
 
 code_reserved SML
--- a/src/HOL/Orderings.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Orderings.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -640,8 +640,8 @@
 let
   val All_binder = Syntax.binder_name @{const_syntax All};
   val Ex_binder = Syntax.binder_name @{const_syntax Ex};
-  val impl = @{const_syntax "op -->"};
-  val conj = @{const_syntax "op &"};
+  val impl = @{const_syntax HOL.implies};
+  val conj = @{const_syntax HOL.conj};
   val less = @{const_syntax less};
   val less_eq = @{const_syntax less_eq};
 
--- a/src/HOL/Predicate.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Predicate.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -808,8 +808,12 @@
 
 lemma eq_pred_code [code]:
   fixes P Q :: "'a pred"
-  shows "eq_class.eq P Q \<longleftrightarrow> P \<le> Q \<and> Q \<le> P"
-  unfolding eq by auto
+  shows "HOL.equal P Q \<longleftrightarrow> P \<le> Q \<and> Q \<le> P"
+  by (auto simp add: equal)
+
+lemma [code nbe]:
+  "HOL.equal (x :: 'a pred) x \<longleftrightarrow> True"
+  by (fact equal_refl)
 
 lemma [code]:
   "pred_case f P = f (eval P)"
--- a/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -4,15 +4,42 @@
 
 section {* Example append *}
 
+
 inductive append
 where
   "append [] ys ys"
 | "append xs ys zs ==> append (x # xs) ys (x # zs)"
 
-code_pred append .
+setup {* Code_Prolog.map_code_options (K
+  {ensure_groundness = false,
+   limited_types = [],
+   limited_predicates = [],
+   replacing = [],
+   manual_reorder = [],
+   prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+values "{(x, y, z). append x y z}"
 
 values 3 "{(x, y, z). append x y z}"
 
+setup {* Code_Prolog.map_code_options (K
+  {ensure_groundness = false,
+   limited_types = [],
+   limited_predicates = [],
+   replacing = [],
+   manual_reorder = [],
+   prolog_system = Code_Prolog.YAP}) *}
+
+values "{(x, y, z). append x y z}"
+
+setup {* Code_Prolog.map_code_options (K
+  {ensure_groundness = false,
+   limited_types = [],
+   limited_predicates = [],
+   replacing = [],
+   manual_reorder = [],
+   prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
 
 section {* Example queens *}
 
@@ -71,9 +98,7 @@
 where
   "queen [1,2,3,4,5,6,7,8,9] ys ==> queen_9 ys"
 
-code_pred queen_9 .
-
-values 150 "{ys. queen_9 ys}"
+values 10 "{ys. queen_9 ys}"
 
 
 section {* Example symbolic derivation *}
@@ -163,14 +188,41 @@
 where
   "d (Mult (Mult (Mult (Mult (Mult (Mult (Mult (Mult (Mult x x) x) x) x) x) x) x) x) x) x e ==> times10 e"
 
-code_pred ops8 .
-code_pred divide10 .
-code_pred log10 .
-code_pred times10 .
-
 values "{e. ops8 e}"
 values "{e. divide10 e}"
 values "{e. log10 e}"
 values "{e. times10 e}"
 
+section {* Example negation *}
+
+datatype abc = A | B | C
+
+inductive notB :: "abc => bool"
+where
+  "y \<noteq> B \<Longrightarrow> notB y"
+
+setup {* Code_Prolog.map_code_options (K
+  {ensure_groundness = true,
+   limited_types = [],
+   limited_predicates = [],
+   replacing = [],
+   manual_reorder = [], 
+   prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+values 2 "{y. notB y}"
+
+inductive notAB :: "abc * abc \<Rightarrow> bool"
+where
+  "y \<noteq> A \<Longrightarrow> z \<noteq> B \<Longrightarrow> notAB (y, z)"
+
+values 5 "{y. notAB y}"
+
+section {* Example prolog conform variable names *}
+
+inductive equals :: "abc => abc => bool"
+where
+  "equals y' y'"
+
+values 1 "{(y, z). equals y z}"
+
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Predicate_Compile_Examples/Context_Free_Grammar_Example.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -0,0 +1,169 @@
+theory Context_Free_Grammar_Example
+imports Code_Prolog
+begin
+
+declare mem_def[code_pred_inline]
+
+
+subsection {* Alternative rules for length *}
+
+definition size_list :: "'a list => nat"
+where "size_list = size"
+
+lemma size_list_simps:
+  "size_list [] = 0"
+  "size_list (x # xs) = Suc (size_list xs)"
+by (auto simp add: size_list_def)
+
+declare size_list_simps[code_pred_def]
+declare size_list_def[symmetric, code_pred_inline]
+
+
+setup {* Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck) *}
+
+datatype alphabet = a | b
+
+inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
+  "[] \<in> S\<^isub>1"
+| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
+| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
+| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
+| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
+| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
+
+lemma
+  "w \<in> S\<^isub>1 \<Longrightarrow> w = []"
+quickcheck[generator = prolog, iterations=1, expect = counterexample]
+oops
+
+definition "filter_a = filter (\<lambda>x. x = a)"
+
+lemma [code_pred_def]: "filter_a [] = []"
+unfolding filter_a_def by simp
+
+lemma [code_pred_def]: "filter_a (x#xs) = (if x = a then x # filter_a xs else filter_a xs)"
+unfolding filter_a_def by simp
+
+declare filter_a_def[symmetric, code_pred_inline]
+
+definition "filter_b = filter (\<lambda>x. x = b)"
+
+lemma [code_pred_def]: "filter_b [] = []"
+unfolding filter_b_def by simp
+
+lemma [code_pred_def]: "filter_b (x#xs) = (if x = b then x # filter_b xs else filter_b xs)"
+unfolding filter_b_def by simp
+
+declare filter_b_def[symmetric, code_pred_inline]
+
+setup {* Code_Prolog.map_code_options (K
+  {ensure_groundness = true,
+  limited_types = [],
+  limited_predicates = [(["s1", "a1", "b1"], 2)],
+  replacing = [(("s1", "limited_s1"), "quickcheck")],
+  manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])],
+  prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+
+theorem S\<^isub>1_sound:
+"w \<in> S\<^isub>1 \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator = prolog, iterations=1, expect = counterexample]
+oops
+
+
+inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
+  "[] \<in> S\<^isub>2"
+| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
+| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
+| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
+| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
+| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
+
+
+setup {* Code_Prolog.map_code_options (K
+  {ensure_groundness = true,
+  limited_types = [],
+  limited_predicates = [(["s2", "a2", "b2"], 3)],
+  replacing = [(("s2", "limited_s2"), "quickcheck")],
+  manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])],
+  prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+
+theorem S\<^isub>2_sound:
+"w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator=prolog, iterations=1, expect = counterexample]
+oops
+
+inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
+  "[] \<in> S\<^isub>3"
+| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
+| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
+| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
+| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
+| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
+
+
+setup {* Code_Prolog.map_code_options (K
+  {ensure_groundness = true,
+  limited_types = [],
+  limited_predicates = [(["s3", "a3", "b3"], 6)],
+  replacing = [(("s3", "limited_s3"), "quickcheck")],
+  manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])],
+  prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+lemma S\<^isub>3_sound:
+"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator=prolog, iterations=1, size=1, expect = no_counterexample]
+oops
+
+
+(*
+setup {* Code_Prolog.map_code_options (K
+  {ensure_groundness = true,
+  limited_types = [],
+  limited_predicates = [],
+  replacing = [],
+  manual_reorder = [],
+  prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+
+theorem S\<^isub>3_complete:
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>3"
+quickcheck[generator = prolog, size=1, iterations=1]
+oops
+*)
+
+inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
+  "[] \<in> S\<^isub>4"
+| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
+| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
+| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
+| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
+| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
+| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
+
+
+setup {* Code_Prolog.map_code_options (K
+  {ensure_groundness = true,
+  limited_types = [],
+  limited_predicates = [(["s4", "a4", "b4"], 6)],
+  replacing = [(("s4", "limited_s4"), "quickcheck")],
+  manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])],
+  prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+
+theorem S\<^isub>4_sound:
+"w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator = prolog, size=1, iterations=1, expect = no_counterexample]
+oops
+
+(*
+theorem S\<^isub>4_complete:
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
+oops
+*)
+
+hide_const a b
+
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Predicate_Compile_Examples/Hotel_Example.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -0,0 +1,129 @@
+theory Hotel_Example
+imports Predicate_Compile_Alternative_Defs Code_Prolog
+begin
+
+datatype guest = Guest0 | Guest1
+datatype key = Key0 | Key1 | Key2 | Key3
+datatype room = Room0
+
+types card = "key * key"
+
+datatype event =
+   Check_in guest room card | Enter guest room card | Exit guest room
+
+definition initk :: "room \<Rightarrow> key"
+  where "initk = (%r. Key0)"
+
+declare initk_def[code_pred_def, code]
+
+primrec owns :: "event list \<Rightarrow> room \<Rightarrow> guest option"
+where
+  "owns [] r = None"
+| "owns (e#s) r = (case e of
+    Check_in g r' c \<Rightarrow> if r' = r then Some g else owns s r |
+    Enter g r' c \<Rightarrow> owns s r |
+    Exit g r' \<Rightarrow> owns s r)"
+
+primrec currk :: "event list \<Rightarrow> room \<Rightarrow> key"
+where
+  "currk [] r = initk r"
+| "currk (e#s) r = (let k = currk s r in
+    case e of Check_in g r' (k1, k2) \<Rightarrow> if r' = r then k2 else k
+            | Enter g r' c \<Rightarrow> k
+            | Exit g r \<Rightarrow> k)"
+
+primrec issued :: "event list \<Rightarrow> key set"
+where
+  "issued [] = range initk"
+| "issued (e#s) = issued s \<union>
+  (case e of Check_in g r (k1, k2) \<Rightarrow> {k2} | Enter g r c \<Rightarrow> {} | Exit g r \<Rightarrow> {})"
+
+primrec cards :: "event list \<Rightarrow> guest \<Rightarrow> card set"
+where
+  "cards [] g = {}"
+| "cards (e#s) g = (let C = cards s g in
+                    case e of Check_in g' r c \<Rightarrow> if g' = g then insert c C
+                                                else C
+                            | Enter g r c \<Rightarrow> C
+                            | Exit g r \<Rightarrow> C)"
+
+primrec roomk :: "event list \<Rightarrow> room \<Rightarrow> key"
+where
+  "roomk [] r = initk r"
+| "roomk (e#s) r = (let k = roomk s r in
+    case e of Check_in g r' c \<Rightarrow> k
+            | Enter g r' (x,y) \<Rightarrow> if r' = r (*& x = k*) then y else k
+            | Exit g r \<Rightarrow> k)"
+
+primrec isin :: "event list \<Rightarrow> room \<Rightarrow> guest set"
+where
+  "isin [] r = {}"
+| "isin (e#s) r = (let G = isin s r in
+                 case e of Check_in g r c \<Rightarrow> G
+                 | Enter g r' c \<Rightarrow> if r' = r then {g} \<union> G else G
+                 | Exit g r' \<Rightarrow> if r'=r then G - {g} else G)"
+
+primrec hotel :: "event list \<Rightarrow> bool"
+where
+  "hotel []  = True"
+| "hotel (e # s) = (hotel s & (case e of
+  Check_in g r (k,k') \<Rightarrow> k = currk s r \<and> k' \<notin> issued s |
+  Enter g r (k,k') \<Rightarrow> (k,k') : cards s g & (roomk s r : {k, k'}) |
+  Exit g r \<Rightarrow> g : isin s r))"
+
+lemma issued_nil: "issued [] = {Key0}"
+by (auto simp add: initk_def)
+
+lemmas issued_simps[code, code_pred_def] = issued_nil issued.simps(2)
+
+declare Let_def[code_pred_inline]
+
+lemma [code_pred_inline]: "insert == (%y A x. y = x | A x)"
+by (auto simp add: insert_iff[unfolded mem_def] expand_fun_eq intro!: eq_reflection)
+
+lemma [code_pred_inline]: "(op -) == (%A B x. A x \<and> \<not> B x)"
+by (auto simp add: Diff_iff[unfolded mem_def] expand_fun_eq intro!: eq_reflection)
+
+setup {* Code_Prolog.map_code_options (K
+  {ensure_groundness = true,
+  limited_types = [],
+  limited_predicates = [],
+  replacing = [],
+  manual_reorder = [],
+  prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+values 40 "{s. hotel s}"
+
+
+setup {* Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck) *}
+
+lemma "\<lbrakk> hotel s; g \<in> isin s r \<rbrakk> \<Longrightarrow> owns s r = Some g"
+quickcheck[generator = code, iterations = 100000, report]
+quickcheck[generator = prolog, iterations = 1, expect = counterexample]
+oops
+
+
+definition no_Check_in :: "event list \<Rightarrow> room \<Rightarrow> bool" where(*>*)
+[code del]: "no_Check_in s r \<equiv> \<not>(\<exists>g c. Check_in g r c \<in> set s)"
+
+
+definition feels_safe :: "event list \<Rightarrow> room \<Rightarrow> bool"
+where
+  "feels_safe s r = (\<exists>s\<^isub>1 s\<^isub>2 s\<^isub>3 g c c'.
+   s = s\<^isub>3 @ [Enter g r c] @ s\<^isub>2 @ [Check_in g r c'] @ s\<^isub>1 \<and>
+   no_Check_in (s\<^isub>3 @ s\<^isub>2) r \<and> isin (s\<^isub>2 @ [Check_in g r c] @ s\<^isub>1) r = {})"
+
+setup {* Code_Prolog.map_code_options (K 
+  {ensure_groundness = true,
+   limited_types = [],
+   limited_predicates = [(["hotel"], 5)],
+   replacing = [(("hotel", "limited_hotel"), "quickcheck")],
+   manual_reorder = [],
+   prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+lemma
+  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
+quickcheck[generator = prolog, iterations = 1, expect = counterexample]
+oops
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Predicate_Compile_Examples/Lambda_Example.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -0,0 +1,126 @@
+theory Lambda_Example
+imports Code_Prolog
+begin
+
+subsection {* Lambda *}
+
+datatype type =
+    Atom nat
+  | Fun type type    (infixr "\<Rightarrow>" 200)
+
+datatype dB =
+    Var nat
+  | App dB dB (infixl "\<degree>" 200)
+  | Abs type dB
+
+primrec
+  nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
+where
+  "[]\<langle>i\<rangle> = None"
+| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
+
+inductive nth_el1 :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
+where
+  "nth_el1 (x # xs) 0 x"
+| "nth_el1 xs i y \<Longrightarrow> nth_el1 (x # xs) (Suc i) y"
+
+inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool"  ("_ \<turnstile> _ : _" [50, 50, 50] 50)
+  where
+    Var [intro!]: "nth_el1 env x T \<Longrightarrow> env \<turnstile> Var x : T"
+  | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
+  | App [intro!]: "env \<turnstile> s : U \<Rightarrow> T \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
+
+primrec
+  lift :: "[dB, nat] => dB"
+where
+    "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
+  | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
+  | "lift (Abs T s) k = Abs T (lift s (k + 1))"
+
+primrec pred :: "nat => nat"
+where
+  "pred (Suc i) = i"
+
+primrec
+  subst :: "[dB, dB, nat] => dB"  ("_[_'/_]" [300, 0, 0] 300)
+where
+    subst_Var: "(Var i)[s/k] =
+      (if k < i then Var (pred i) else if i = k then s else Var i)"
+  | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
+  | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
+
+inductive beta :: "[dB, dB] => bool"  (infixl "\<rightarrow>\<^sub>\<beta>" 50)
+  where
+    beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
+  | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
+  | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
+  | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
+
+subsection {* Inductive definitions for ordering on naturals *}
+
+inductive less_nat
+where
+  "less_nat 0 (Suc y)"
+| "less_nat x y ==> less_nat (Suc x) (Suc y)"
+
+lemma less_nat[code_pred_inline]:
+  "x < y = less_nat x y"
+apply (rule iffI)
+apply (induct x arbitrary: y)
+apply (case_tac y) apply (auto intro: less_nat.intros)
+apply (case_tac y)
+apply (auto intro: less_nat.intros)
+apply (induct rule: less_nat.induct)
+apply auto
+done
+
+lemma [code_pred_inline]: "(x::nat) + 1 = Suc x"
+by simp
+
+section {* Manual setup to find counterexample *}
+
+setup {* Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck) *}
+
+setup {* Code_Prolog.map_code_options (K 
+  { ensure_groundness = true,
+    limited_types = [(@{typ nat}, 1), (@{typ "type"}, 1), (@{typ dB}, 1), (@{typ "type list"}, 1)],
+    limited_predicates = [(["typing"], 2), (["nthel1"], 2)],
+    replacing = [(("typing", "limited_typing"), "quickcheck"),
+                 (("nthel1", "limited_nthel1"), "lim_typing")],
+    manual_reorder = [],
+    prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+lemma
+  "\<Gamma> \<turnstile> t : U \<Longrightarrow> t \<rightarrow>\<^sub>\<beta> t' \<Longrightarrow> \<Gamma> \<turnstile> t' : U"
+quickcheck[generator = prolog, iterations = 1, expect = counterexample]
+oops
+
+text {* Verifying that the found counterexample really is one by means of a proof *}
+
+lemma
+assumes
+  "t' = Var 0"
+  "U = Atom 0"
+  "\<Gamma> = [Atom 1]"
+  "t = App (Abs (Atom 0) (Var 1)) (Var 0)"
+shows
+  "\<Gamma> \<turnstile> t : U"
+  "t \<rightarrow>\<^sub>\<beta> t'"
+  "\<not> \<Gamma> \<turnstile> t' : U"
+proof -
+  from assms show "\<Gamma> \<turnstile> t : U"
+    by (auto intro!: typing.intros nth_el1.intros)
+next
+  from assms have "t \<rightarrow>\<^sub>\<beta> (Var 1)[Var 0/0]"
+    by (auto simp only: intro: beta.intros)
+  moreover
+  from assms have "(Var 1)[Var 0/0] = t'" by simp
+  ultimately show "t \<rightarrow>\<^sub>\<beta> t'" by simp
+next
+  from assms show "\<not> \<Gamma> \<turnstile> t' : U"
+    by (auto elim: typing.cases nth_el1.cases)
+qed
+
+
+end
+
--- a/src/HOL/Predicate_Compile_Examples/ROOT.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Predicate_Compile_Examples/ROOT.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -1,2 +1,2 @@
 use_thys ["Predicate_Compile_Examples", "Predicate_Compile_Quickcheck_Examples", "Specialisation_Examples"];
-if getenv "PROLOG_HOME" = "" then () else use_thys ["Code_Prolog_Examples"];
+if getenv "PROLOG_HOME" = "" then () else use_thys ["Code_Prolog_Examples", "Hotel_Example", "Lambda_Example"];
--- a/src/HOL/Product_Type.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Product_Type.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -21,17 +21,17 @@
   -- "prefer plain propositional version"
 
 lemma
-  shows [code]: "eq_class.eq False P \<longleftrightarrow> \<not> P"
-    and [code]: "eq_class.eq True P \<longleftrightarrow> P" 
-    and [code]: "eq_class.eq P False \<longleftrightarrow> \<not> P" 
-    and [code]: "eq_class.eq P True \<longleftrightarrow> P"
-    and [code nbe]: "eq_class.eq P P \<longleftrightarrow> True"
-  by (simp_all add: eq)
+  shows [code]: "HOL.equal False P \<longleftrightarrow> \<not> P"
+    and [code]: "HOL.equal True P \<longleftrightarrow> P" 
+    and [code]: "HOL.equal P False \<longleftrightarrow> \<not> P" 
+    and [code]: "HOL.equal P True \<longleftrightarrow> P"
+    and [code nbe]: "HOL.equal P P \<longleftrightarrow> True"
+  by (simp_all add: equal)
 
-code_const "eq_class.eq \<Colon> bool \<Rightarrow> bool \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> bool \<Rightarrow> bool \<Rightarrow> bool"
   (Haskell infixl 4 "==")
 
-code_instance bool :: eq
+code_instance bool :: equal
   (Haskell -)
 
 
@@ -92,7 +92,7 @@
 end
 
 lemma [code]:
-  "eq_class.eq (u\<Colon>unit) v \<longleftrightarrow> True" unfolding eq unit_eq [of u] unit_eq [of v] by rule+
+  "HOL.equal (u\<Colon>unit) v \<longleftrightarrow> True" unfolding equal unit_eq [of u] unit_eq [of v] by rule+
 
 code_type unit
   (SML "unit")
@@ -106,10 +106,10 @@
   (Haskell "()")
   (Scala "()")
 
-code_instance unit :: eq
+code_instance unit :: equal
   (Haskell -)
 
-code_const "eq_class.eq \<Colon> unit \<Rightarrow> unit \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> unit \<Rightarrow> unit \<Rightarrow> bool"
   (Haskell infixl 4 "==")
 
 code_reserved SML
@@ -277,10 +277,10 @@
   (Haskell "!((_),/ (_))")
   (Scala "!((_),/ (_))")
 
-code_instance prod :: eq
+code_instance prod :: equal
   (Haskell -)
 
-code_const "eq_class.eq \<Colon> 'a\<Colon>eq \<times> 'b\<Colon>eq \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> 'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool"
   (Haskell infixl 4 "==")
 
 types_code
--- a/src/HOL/Prolog/prolog.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Prolog/prolog.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -2,7 +2,7 @@
     Author:   David von Oheimb (based on a lecture on Lambda Prolog by Nadathur)
 *)
 
-Unsynchronized.set Proof.show_main_goal;
+Proof.show_main_goal := true;
 
 structure Prolog =
 struct
@@ -11,12 +11,12 @@
 
 fun isD t = case t of
     Const(@{const_name Trueprop},_)$t     => isD t
-  | Const(@{const_name "op &"}  ,_)$l$r     => isD l andalso isD r
-  | Const(@{const_name "op -->"},_)$l$r     => isG l andalso isD r
+  | Const(@{const_name HOL.conj}  ,_)$l$r     => isD l andalso isD r
+  | Const(@{const_name HOL.implies},_)$l$r     => isG l andalso isD r
   | Const(   "==>",_)$l$r     => isG l andalso isD r
   | Const(@{const_name All},_)$Abs(s,_,t) => isD t
   | Const("all",_)$Abs(s,_,t) => isD t
-  | Const(@{const_name "op |"},_)$_$_       => false
+  | Const(@{const_name HOL.disj},_)$_$_       => false
   | Const(@{const_name Ex} ,_)$_          => false
   | Const(@{const_name Not},_)$_          => false
   | Const(@{const_name True},_)           => false
@@ -30,9 +30,9 @@
 and
     isG t = case t of
     Const(@{const_name Trueprop},_)$t     => isG t
-  | Const(@{const_name "op &"}  ,_)$l$r     => isG l andalso isG r
-  | Const(@{const_name "op |"}  ,_)$l$r     => isG l andalso isG r
-  | Const(@{const_name "op -->"},_)$l$r     => isD l andalso isG r
+  | Const(@{const_name HOL.conj}  ,_)$l$r     => isG l andalso isG r
+  | Const(@{const_name HOL.disj}  ,_)$l$r     => isG l andalso isG r
+  | Const(@{const_name HOL.implies},_)$l$r     => isD l andalso isG r
   | Const(   "==>",_)$l$r     => isD l andalso isG r
   | Const(@{const_name All},_)$Abs(_,_,t) => isG t
   | Const("all",_)$Abs(_,_,t) => isG t
@@ -53,8 +53,8 @@
     fun at  thm = case concl_of thm of
       _$(Const(@{const_name All} ,_)$Abs(s,_,_))=> at(thm RS
         (read_instantiate ctxt [(("x", 0), "?" ^ (if s="P" then "PP" else s))] spec))
-    | _$(Const(@{const_name "op &"},_)$_$_)       => at(thm RS conjunct1)@at(thm RS conjunct2)
-    | _$(Const(@{const_name "op -->"},_)$_$_)     => at(thm RS mp)
+    | _$(Const(@{const_name HOL.conj},_)$_$_)       => at(thm RS conjunct1)@at(thm RS conjunct2)
+    | _$(Const(@{const_name HOL.implies},_)$_$_)     => at(thm RS mp)
     | _                             => [thm]
 in map zero_var_indexes (at thm) end;
 
@@ -72,7 +72,7 @@
   -- is nice, but cannot instantiate unknowns in the assumptions *)
 fun hyp_resolve_tac i st = let
         fun ap (Const(@{const_name All},_)$Abs(_,_,t))=(case ap t of (k,a,t) => (k+1,a  ,t))
-        |   ap (Const(@{const_name "op -->"},_)$_$t)    =(case ap t of (k,_,t) => (k,true ,t))
+        |   ap (Const(@{const_name HOL.implies},_)$_$t)    =(case ap t of (k,_,t) => (k,true ,t))
         |   ap t                          =                         (0,false,t);
 (*
         fun rep_goal (Const ("all",_)$Abs (_,_,t)) = rep_goal t
--- a/src/HOL/Quickcheck.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Quickcheck.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -97,7 +97,7 @@
   \<Rightarrow> Random.seed \<Rightarrow> (('a\<Colon>term_of \<Rightarrow> 'b\<Colon>typerep) \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
   "random_fun_lift f = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Evaluation.term_of f Random.split_seed"
 
-instantiation "fun" :: ("{eq, term_of}", random) random
+instantiation "fun" :: ("{equal, term_of}", random) random
 begin
 
 definition random_fun :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
--- a/src/HOL/Quotient.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Quotient.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -651,6 +651,16 @@
   shows "(R1 ===> (R1 ===> R2) ===> R2) Let Let"
   by auto
 
+lemma mem_rsp:
+  shows "(R1 ===> (R1 ===> R2) ===> R2) op \<in> op \<in>"
+  by (simp add: mem_def)
+
+lemma mem_prs:
+  assumes a1: "Quotient R1 Abs1 Rep1"
+  and     a2: "Quotient R2 Abs2 Rep2"
+  shows "(Rep1 ---> (Abs1 ---> Rep2) ---> Abs2) op \<in> = op \<in>"
+  by (simp add: expand_fun_eq mem_def Quotient_abs_rep[OF a1] Quotient_abs_rep[OF a2])
+
 locale quot_type =
   fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
   and   Abs :: "('a \<Rightarrow> bool) \<Rightarrow> 'b"
@@ -721,8 +731,8 @@
 declare [[map "fun" = (fun_map, fun_rel)]]
 
 lemmas [quot_thm] = fun_quotient
-lemmas [quot_respect] = quot_rel_rsp if_rsp o_rsp let_rsp
-lemmas [quot_preserve] = if_prs o_prs let_prs
+lemmas [quot_respect] = quot_rel_rsp if_rsp o_rsp let_rsp mem_rsp
+lemmas [quot_preserve] = if_prs o_prs let_prs mem_prs
 lemmas [quot_equiv] = identity_equivp
 
 
@@ -773,20 +783,20 @@
 
 method_setup lifting =
   {* Attrib.thms >> (fn thms => fn ctxt => 
-       SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.lift_tac ctxt thms))) *}
+       SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.lift_tac ctxt [] thms))) *}
   {* lifts theorems to quotient types *}
 
 method_setup lifting_setup =
   {* Attrib.thm >> (fn thm => fn ctxt => 
-       SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.lift_procedure_tac ctxt thm))) *}
+       SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.lift_procedure_tac ctxt [] thm))) *}
   {* sets up the three goals for the quotient lifting procedure *}
 
 method_setup descending =
-  {* Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.descend_tac ctxt))) *}
+  {* Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.descend_tac ctxt []))) *}
   {* decends theorems to the raw level *}
 
 method_setup descending_setup =
-  {* Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.descend_procedure_tac ctxt))) *}
+  {* Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.descend_procedure_tac ctxt []))) *}
   {* sets up the three goals for the decending theorems *}
 
 method_setup regularize =
--- a/src/HOL/Rat.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Rat.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -1083,18 +1083,18 @@
   by (cases "0::int" a rule: linorder_cases) (simp_all add: quotient_of_Fract)
 qed
 
-instantiation rat :: eq
+instantiation rat :: equal
 begin
 
 definition [code]:
-  "eq_class.eq a b \<longleftrightarrow> quotient_of a = quotient_of b"
+  "HOL.equal a b \<longleftrightarrow> quotient_of a = quotient_of b"
 
 instance proof
-qed (simp add: eq_rat_def quotient_of_inject_eq)
+qed (simp add: equal_rat_def quotient_of_inject_eq)
 
 lemma rat_eq_refl [code nbe]:
-  "eq_class.eq (r::rat) r \<longleftrightarrow> True"
-  by (rule HOL.eq_refl)
+  "HOL.equal (r::rat) r \<longleftrightarrow> True"
+  by (rule equal_refl)
 
 end
 
--- a/src/HOL/RealDef.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/RealDef.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -1697,19 +1697,21 @@
   "Ratreal (number_of r / number_of s) = number_of r / number_of s"
 unfolding Ratreal_number_of_quotient [symmetric] Ratreal_def of_rat_divide ..
 
-instantiation real :: eq
+instantiation real :: equal
 begin
 
-definition "eq_class.eq (x\<Colon>real) y \<longleftrightarrow> x - y = 0"
+definition "HOL.equal (x\<Colon>real) y \<longleftrightarrow> x - y = 0"
 
-instance by default (simp add: eq_real_def)
+instance proof
+qed (simp add: equal_real_def)
 
-lemma real_eq_code [code]: "eq_class.eq (Ratreal x) (Ratreal y) \<longleftrightarrow> eq_class.eq x y"
-  by (simp add: eq_real_def eq)
+lemma real_equal_code [code]:
+  "HOL.equal (Ratreal x) (Ratreal y) \<longleftrightarrow> HOL.equal x y"
+  by (simp add: equal_real_def equal)
 
-lemma real_eq_refl [code nbe]:
-  "eq_class.eq (x::real) x \<longleftrightarrow> True"
-  by (rule HOL.eq_refl)
+lemma [code nbe]:
+  "HOL.equal (x::real) x \<longleftrightarrow> True"
+  by (rule equal_refl)
 
 end
 
--- a/src/HOL/Set.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Set.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -219,8 +219,8 @@
   val Type (set_type, _) = @{typ "'a set"};   (* FIXME 'a => bool (!?!) *)
   val All_binder = Syntax.binder_name @{const_syntax All};
   val Ex_binder = Syntax.binder_name @{const_syntax Ex};
-  val impl = @{const_syntax "op -->"};
-  val conj = @{const_syntax "op &"};
+  val impl = @{const_syntax HOL.implies};
+  val conj = @{const_syntax HOL.conj};
   val sbset = @{const_syntax subset};
   val sbset_eq = @{const_syntax subset_eq};
 
@@ -268,8 +268,8 @@
 
     fun setcompr_tr [e, idts, b] =
       let
-        val eq = Syntax.const @{const_syntax "op ="} $ Bound (nvars idts) $ e;
-        val P = Syntax.const @{const_syntax "op &"} $ eq $ b;
+        val eq = Syntax.const @{const_syntax HOL.eq} $ Bound (nvars idts) $ e;
+        val P = Syntax.const @{const_syntax HOL.conj} $ eq $ b;
         val exP = ex_tr [idts, P];
       in Syntax.const @{const_syntax Collect} $ Term.absdummy (dummyT, exP) end;
 
@@ -288,8 +288,8 @@
   fun setcompr_tr' [Abs (abs as (_, _, P))] =
     let
       fun check (Const (@{const_syntax Ex}, _) $ Abs (_, _, P), n) = check (P, n + 1)
-        | check (Const (@{const_syntax "op &"}, _) $
-              (Const (@{const_syntax "op ="}, _) $ Bound m $ e) $ P, n) =
+        | check (Const (@{const_syntax HOL.conj}, _) $
+              (Const (@{const_syntax HOL.eq}, _) $ Bound m $ e) $ P, n) =
             n > 0 andalso m = n andalso not (loose_bvar1 (P, n)) andalso
             subset (op =) (0 upto (n - 1), add_loose_bnos (e, 0, []))
         | check _ = false;
@@ -305,7 +305,7 @@
           val M = Syntax.const @{syntax_const "_Coll"} $ x $ t;
         in
           case t of
-            Const (@{const_syntax "op &"}, _) $
+            Const (@{const_syntax HOL.conj}, _) $
               (Const (@{const_syntax Set.member}, _) $
                 (Const (@{syntax_const "_bound"}, _) $ Free (yN, _)) $ A) $ P =>
             if xN = yN then Syntax.const @{syntax_const "_Collect"} $ x $ A $ P else M
--- a/src/HOL/Statespace/DistinctTreeProver.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/DistinctTreeProver.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -32,17 +32,18 @@
 subsection {* Distinctness of Nodes *}
 
 
-consts set_of:: "'a tree \<Rightarrow> 'a set"
-primrec 
-"set_of Tip = {}"
-"set_of (Node l x d r) = (if d then {} else {x}) \<union> set_of l \<union> set_of r"
+primrec set_of :: "'a tree \<Rightarrow> 'a set"
+where
+  "set_of Tip = {}"
+| "set_of (Node l x d r) = (if d then {} else {x}) \<union> set_of l \<union> set_of r"
 
-consts all_distinct:: "'a tree \<Rightarrow> bool"
-primrec
-"all_distinct Tip = True"
-"all_distinct (Node l x d r) = ((d \<or> (x \<notin> set_of l \<and> x \<notin> set_of r)) \<and> 
-                               set_of l \<inter> set_of r = {} \<and>
-                               all_distinct l \<and> all_distinct r)"
+primrec all_distinct :: "'a tree \<Rightarrow> bool"
+where
+  "all_distinct Tip = True"
+| "all_distinct (Node l x d r) =
+    ((d \<or> (x \<notin> set_of l \<and> x \<notin> set_of r)) \<and> 
+      set_of l \<inter> set_of r = {} \<and>
+      all_distinct l \<and> all_distinct r)"
 
 text {* Given a binary tree @{term "t"} for which 
 @{const all_distinct} holds, given two different nodes contained in the tree,
@@ -99,19 +100,19 @@
 *}
 
 
-consts "delete" :: "'a \<Rightarrow> 'a tree \<Rightarrow> 'a tree option"
-primrec
-"delete x Tip = None"
-"delete x (Node l y d r) = (case delete x l of
-                              Some l' \<Rightarrow>
-                               (case delete x r of 
-                                  Some r' \<Rightarrow> Some (Node l' y (d \<or> (x=y)) r')
-                                | None \<Rightarrow> Some (Node l' y (d \<or> (x=y)) r))
-                             | None \<Rightarrow>
-                                (case (delete x r) of 
-                                   Some r' \<Rightarrow> Some (Node l y (d \<or> (x=y)) r')
-                                 | None \<Rightarrow> if x=y \<and> \<not>d then Some (Node l y True r)
-                                           else None))"
+primrec delete :: "'a \<Rightarrow> 'a tree \<Rightarrow> 'a tree option"
+where
+  "delete x Tip = None"
+| "delete x (Node l y d r) = (case delete x l of
+                                Some l' \<Rightarrow>
+                                 (case delete x r of 
+                                    Some r' \<Rightarrow> Some (Node l' y (d \<or> (x=y)) r')
+                                  | None \<Rightarrow> Some (Node l' y (d \<or> (x=y)) r))
+                               | None \<Rightarrow>
+                                  (case (delete x r) of 
+                                     Some r' \<Rightarrow> Some (Node l y (d \<or> (x=y)) r')
+                                   | None \<Rightarrow> if x=y \<and> \<not>d then Some (Node l y True r)
+                                             else None))"
 
 
 lemma delete_Some_set_of: "\<And>t'. delete x t = Some t' \<Longrightarrow> set_of t' \<subseteq> set_of t"
@@ -293,15 +294,15 @@
 qed
 
 
-consts "subtract" :: "'a tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree option"
-primrec
-"subtract Tip t = Some t"
-"subtract (Node l x b r) t = 
-   (case delete x t of
-      Some t' \<Rightarrow> (case subtract l t' of 
-                   Some t'' \<Rightarrow> subtract r t''
-                  | None \<Rightarrow> None)
-    | None \<Rightarrow> None)"
+primrec subtract :: "'a tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree option"
+where
+  "subtract Tip t = Some t"
+| "subtract (Node l x b r) t =
+     (case delete x t of
+        Some t' \<Rightarrow> (case subtract l t' of 
+                     Some t'' \<Rightarrow> subtract r t''
+                    | None \<Rightarrow> None)
+       | None \<Rightarrow> None)"
 
 lemma subtract_Some_set_of_res: 
   "\<And>t\<^isub>2 t. subtract t\<^isub>1 t\<^isub>2 = Some t \<Longrightarrow> set_of t \<subseteq> set_of t\<^isub>2"
--- a/src/HOL/Statespace/StateFun.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/StateFun.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -33,10 +33,10 @@
 lemma K_statefun_cong [cong]: "K_statefun c x = K_statefun c x"
   by (rule refl)
 
-definition lookup:: "('v \<Rightarrow> 'a) \<Rightarrow> 'n \<Rightarrow> ('n \<Rightarrow> 'v) \<Rightarrow> 'a"
+definition lookup :: "('v \<Rightarrow> 'a) \<Rightarrow> 'n \<Rightarrow> ('n \<Rightarrow> 'v) \<Rightarrow> 'a"
   where "lookup destr n s = destr (s n)"
 
-definition update::
+definition update ::
   "('v \<Rightarrow> 'a1) \<Rightarrow> ('a2 \<Rightarrow> 'v) \<Rightarrow> 'n \<Rightarrow> ('a1 \<Rightarrow> 'a2) \<Rightarrow> ('n \<Rightarrow> 'v) \<Rightarrow> ('n \<Rightarrow> 'v)"
   where "update destr constr n f s = s(n := constr (f (destr (s n))))"
 
--- a/src/HOL/Statespace/StateSpaceEx.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/StateSpaceEx.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -1,14 +1,12 @@
 (*  Title:      StateSpaceEx.thy
-    ID:         $Id$
     Author:     Norbert Schirmer, TU Muenchen
 *)
 
 header {* Examples \label{sec:Examples} *}
 theory StateSpaceEx
 imports StateSpaceLocale StateSpaceSyntax
+begin
 
-begin
-(* FIXME: Use proper keywords file *)
 (*<*)
 syntax
  "_statespace_updates" :: "('a \<Rightarrow> 'b) \<Rightarrow> updbinds \<Rightarrow> ('a \<Rightarrow> 'b)" ("_\<langle>_\<rangle>" [900,0] 900)
--- a/src/HOL/Statespace/StateSpaceLocale.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/StateSpaceLocale.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -1,5 +1,4 @@
 (*  Title:      StateSpaceLocale.thy
-    ID:         $Id$
     Author:     Norbert Schirmer, TU Muenchen
 *)
 
@@ -18,7 +17,7 @@
 
 locale project_inject =
  fixes project :: "'value \<Rightarrow> 'a"
- and   "inject":: "'a \<Rightarrow> 'value"
+  and inject :: "'a \<Rightarrow> 'value"
  assumes project_inject_cancel [statefun_simp]: "project (inject x) = x"
 
 lemma (in project_inject)
--- a/src/HOL/Statespace/StateSpaceSyntax.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/StateSpaceSyntax.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -5,7 +5,6 @@
 header {* Syntax for State Space Lookup and Update \label{sec:StateSpaceSyntax}*}
 theory StateSpaceSyntax
 imports StateSpaceLocale
-
 begin
 
 text {* The state space syntax is kept in an extra theory so that you
--- a/src/HOL/Statespace/distinct_tree_prover.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/distinct_tree_prover.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -343,7 +343,7 @@
   end handle Option => NONE)
 
 fun distinctTree_tac names ctxt 
-      (Const (@{const_name Trueprop},_) $ (Const (@{const_name Not}, _) $ (Const (@{const_name "op ="}, _) $ x $ y)), i) =
+      (Const (@{const_name Trueprop},_) $ (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.eq}, _) $ x $ y)), i) =
   (case get_fst_success (neq_x_y ctxt x y) names of
      SOME neq => rtac neq i
    | NONE => no_tac)
@@ -356,7 +356,7 @@
 
 fun distinct_simproc names =
   Simplifier.simproc_global @{theory HOL} "DistinctTreeProver.distinct_simproc" ["x = y"]
-    (fn thy => fn ss => fn (Const (@{const_name "op ="},_)$x$y) =>
+    (fn thy => fn ss => fn (Const (@{const_name HOL.eq},_)$x$y) =>
         case try Simplifier.the_context ss of
         SOME ctxt => Option.map (fn neq => neq_to_eq_False OF [neq]) 
                       (get_fst_success (neq_x_y ctxt x y) names)
--- a/src/HOL/Statespace/state_fun.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/state_fun.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -53,7 +53,7 @@
 val lazy_conj_simproc =
   Simplifier.simproc_global @{theory HOL} "lazy_conj_simp" ["P & Q"]
     (fn thy => fn ss => fn t =>
-      (case t of (Const (@{const_name "op &"},_)$P$Q) => 
+      (case t of (Const (@{const_name HOL.conj},_)$P$Q) => 
          let
             val P_P' = Simplifier.rewrite ss (cterm_of thy P);
             val P' = P_P' |> prop_of |> Logic.dest_equals |> #2 
@@ -285,7 +285,7 @@
                             then Bound 2
                             else raise TERM ("",[n]);
                    val sel' = lo $ d $ n' $ s;
-                  in (Const (@{const_name "op ="},Teq)$sel'$x',hd (binder_types Teq),nT,swap) end;
+                  in (Const (@{const_name HOL.eq},Teq)$sel'$x',hd (binder_types Teq),nT,swap) end;
 
          fun dest_state (s as Bound 0) = s
            | dest_state (s as (Const (sel,sT)$Bound 0)) =
@@ -295,10 +295,10 @@
            | dest_state s = 
                     raise TERM ("StateFun.ex_lookup_eq_simproc: not a record slector",[s]);
   
-         fun dest_sel_eq (Const (@{const_name "op ="},Teq)$
+         fun dest_sel_eq (Const (@{const_name HOL.eq},Teq)$
                            ((lo as (Const ("StateFun.lookup",lT)))$d$n$s)$X) =
                            (false,Teq,lT,lo,d,n,X,dest_state s)
-           | dest_sel_eq (Const (@{const_name "op ="},Teq)$X$
+           | dest_sel_eq (Const (@{const_name HOL.eq},Teq)$X$
                             ((lo as (Const ("StateFun.lookup",lT)))$d$n$s)) =
                            (true,Teq,lT,lo,d,n,X,dest_state s)
            | dest_sel_eq _ = raise TERM ("",[]);
--- a/src/HOL/Statespace/state_space.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/state_space.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -223,7 +223,7 @@
 
 fun distinctTree_tac ctxt
       (Const (@{const_name Trueprop},_) $
-        (Const (@{const_name Not}, _) $ (Const (@{const_name "op ="}, _) $ (x as Free _)$ (y as Free _))), i) =
+        (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.eq}, _) $ (x as Free _)$ (y as Free _))), i) =
   (case (neq_x_y ctxt x y) of
      SOME neq => rtac neq i
    | NONE => no_tac)
@@ -236,7 +236,7 @@
 
 val distinct_simproc =
   Simplifier.simproc_global @{theory HOL} "StateSpace.distinct_simproc" ["x = y"]
-    (fn thy => fn ss => (fn (Const (@{const_name "op ="},_)$(x as Free _)$(y as Free _)) =>
+    (fn thy => fn ss => (fn (Const (@{const_name HOL.eq},_)$(x as Free _)$(y as Free _)) =>
         (case try Simplifier.the_context ss of
           SOME ctxt => Option.map (fn neq => DistinctTreeProver.neq_to_eq_False OF [neq])
                        (neq_x_y ctxt x y)
@@ -277,28 +277,29 @@
     fun comps_of_thm thm = prop_of thm
              |> (fn (_$(_$t)) => DistinctTreeProver.dest_tree t) |> map (fst o dest_Free);
 
-    fun type_attr phi (ctxt,thm) =
-      (case ctxt of Context.Theory _ => (ctxt,thm)
-       | _ =>
+    fun type_attr phi = Thm.declaration_attribute (fn thm => fn context =>
+      (case context of
+        Context.Theory _ => context
+      | Context.Proof ctxt =>
         let
-          val {declinfo,distinctthm=tt,silent} = (NameSpaceData.get ctxt);
+          val {declinfo,distinctthm=tt,silent} = NameSpaceData.get context;
           val all_names = comps_of_thm thm;
           fun upd name tt =
-               (case (Symtab.lookup tt name) of
+               (case Symtab.lookup tt name of
                  SOME dthm => if sorted_subset (op =) (comps_of_thm dthm) all_names
                               then Symtab.update (name,thm) tt else tt
-                | NONE => Symtab.update (name,thm) tt)
+               | NONE => Symtab.update (name,thm) tt)
 
           val tt' = tt |> fold upd all_names;
           val activate_simproc =
-              Output.no_warnings_CRITICAL   (* FIXME !?! *)
-               (Simplifier.map_ss (fn ss => ss addsimprocs [distinct_simproc]));
-          val ctxt' =
-              ctxt
+            Simplifier.map_ss
+              (Simplifier.with_context (Context_Position.set_visible false ctxt)
+                (fn ss => ss addsimprocs [distinct_simproc]));
+          val context' =
+              context
               |> NameSpaceData.put {declinfo=declinfo,distinctthm=tt',silent=silent}
-              |> activate_simproc
-        in (ctxt',thm)
-        end)
+              |> activate_simproc;
+        in context' end));
 
     val attr = Attrib.internal type_attr;
 
--- a/src/HOL/String.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/String.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -53,7 +53,7 @@
    (fn n => fn m => Drule.instantiate' [] [SOME n, SOME m] @{thm nibble_pair_of_char.simps})
       nibbles nibbles;
 in
-  PureThy.note_thmss Thm.definitionK [((Binding.name "nibble_pair_of_char_simps", []), [(thms, [])])]
+  PureThy.note_thmss Thm.definitionK [((@{binding nibble_pair_of_char_simps}, []), [(thms, [])])]
   #-> (fn [(_, thms)] => fold_rev Code.add_eqn thms)
 end
 *}
@@ -183,10 +183,10 @@
   fold String_Code.add_literal_string ["SML", "OCaml", "Haskell", "Scala"]
 *}
 
-code_instance literal :: eq
+code_instance literal :: equal
   (Haskell -)
 
-code_const "eq_class.eq \<Colon> literal \<Rightarrow> literal \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> literal \<Rightarrow> literal \<Rightarrow> bool"
   (SML "!((_ : string) = _)")
   (OCaml "!((_ : string) = _)")
   (Haskell infixl 4 "==")
--- a/src/HOL/TLA/Intensional.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/TLA/Intensional.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -279,7 +279,7 @@
 
     fun hflatten t =
         case (concl_of t) of
-          Const _ $ (Const (@{const_name "op -->"}, _) $ _ $ _) => hflatten (t RS mp)
+          Const _ $ (Const (@{const_name HOL.implies}, _) $ _ $ _) => hflatten (t RS mp)
         | _ => (hflatten (matchsome conjI t)) handle THM _ => zero_var_indexes t
   in
     hflatten t
--- a/src/HOL/Tools/ATP/atp_systems.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/ATP/atp_systems.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -19,7 +19,7 @@
      has_incomplete_mode: bool,
      proof_delims: (string * string) list,
      known_failures: (failure * string) list,
-     default_max_relevant_per_iter: int,
+     default_max_relevant: int,
      default_theory_relevant: bool,
      explicit_forall: bool,
      use_conjecture_for_hypotheses: bool}
@@ -52,7 +52,7 @@
    has_incomplete_mode: bool,
    proof_delims: (string * string) list,
    known_failures: (failure * string) list,
-   default_max_relevant_per_iter: int,
+   default_max_relevant: int,
    default_theory_relevant: bool,
    explicit_forall: bool,
    use_conjecture_for_hypotheses: bool}
@@ -125,10 +125,20 @@
   priority ("Available ATPs: " ^
             commas (sort_strings (Symtab.keys (Data.get thy))) ^ ".")
 
-fun to_generous_secs time = (Time.toMilliseconds time + 999) div 1000
+fun to_secs bonus time = (Time.toMilliseconds time + bonus + 999) div 1000
 
 (* E prover *)
 
+(* Give older versions of E an extra second, because the "eproof" script wrongly
+   subtracted an entire second to account for the overhead of the script
+   itself, which is in fact much lower. *)
+fun e_bonus () =
+  case getenv "E_VERSION" of
+    "" => 1000
+  | version =>
+    if exists (fn s => String.isPrefix s version) ["0.9", "1.0"] then 1000
+    else 0
+
 val tstp_proof_delims =
   ("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation")
 
@@ -137,8 +147,7 @@
    required_execs = [],
    arguments = fn _ => fn timeout =>
      "--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent \
-     \--soft-cpu-limit=" ^
-     string_of_int (to_generous_secs timeout),
+     \--cpu-limit=" ^ string_of_int (to_secs (e_bonus ()) timeout),
    has_incomplete_mode = false,
    proof_delims = [tstp_proof_delims],
    known_failures =
@@ -150,7 +159,7 @@
        "# Cannot determine problem status within resource limit"),
       (OutOfResources, "SZS status: ResourceOut"),
       (OutOfResources, "SZS status ResourceOut")],
-   default_max_relevant_per_iter = 80 (* FUDGE *),
+   default_max_relevant = 500 (* FUDGE *),
    default_theory_relevant = false,
    explicit_forall = false,
    use_conjecture_for_hypotheses = true}
@@ -165,7 +174,7 @@
    required_execs = [("SPASS_HOME", "SPASS")],
    arguments = fn complete => fn timeout =>
      ("-Auto -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \
-      \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_generous_secs timeout))
+      \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_secs 0 timeout))
      |> not complete ? prefix "-SOS=1 ",
    has_incomplete_mode = true,
    proof_delims = [("Here is a proof", "Formulae used in the proof")],
@@ -177,7 +186,7 @@
       (MalformedInput, "Undefined symbol"),
       (MalformedInput, "Free Variable"),
       (SpassTooOld, "tptp2dfg")],
-   default_max_relevant_per_iter = 40 (* FUDGE *),
+   default_max_relevant = 350 (* FUDGE *),
    default_theory_relevant = true,
    explicit_forall = true,
    use_conjecture_for_hypotheses = true}
@@ -190,10 +199,11 @@
 val vampire_config : prover_config =
   {exec = ("VAMPIRE_HOME", "vampire"),
    required_execs = [],
-   arguments = fn _ => fn timeout =>
-     "--mode casc -t " ^ string_of_int (to_generous_secs timeout) ^
-     " --thanks Andrei --input_file",
-   has_incomplete_mode = false,
+   arguments = fn complete => fn timeout =>
+     ("--mode casc -t " ^ string_of_int (to_secs 0 timeout) ^
+      " --thanks Andrei --input_file")
+     |> not complete ? prefix "--sos on ",
+   has_incomplete_mode = true,
    proof_delims =
      [("=========== Refutation ==========",
        "======= End of refutation ======="),
@@ -206,7 +216,7 @@
       (Unprovable, "Satisfiability detected"),
       (Unprovable, "Termination reason: Satisfiable"),
       (VampireTooOld, "not a valid option")],
-   default_max_relevant_per_iter = 45 (* FUDGE *),
+   default_max_relevant = 400 (* FUDGE *),
    default_theory_relevant = false,
    explicit_forall = false,
    use_conjecture_for_hypotheses = true}
@@ -246,50 +256,50 @@
   | SOME sys => sys
 
 fun remote_config system_name system_versions proof_delims known_failures
-                  default_max_relevant_per_iter default_theory_relevant
-                  use_conjecture_for_hypotheses =
+                  default_max_relevant default_theory_relevant
+                  use_conjecture_for_hypotheses : prover_config =
   {exec = ("ISABELLE_ATP", "scripts/remote_atp"),
    required_execs = [],
    arguments = fn _ => fn timeout =>
-     " -t " ^ string_of_int (to_generous_secs timeout) ^ " -s " ^
+     " -t " ^ string_of_int (to_secs 0 timeout) ^ " -s " ^
      the_system system_name system_versions,
    has_incomplete_mode = false,
    proof_delims = insert (op =) tstp_proof_delims proof_delims,
    known_failures =
      known_failures @ known_perl_failures @
      [(TimedOut, "says Timeout")],
-   default_max_relevant_per_iter = default_max_relevant_per_iter,
+   default_max_relevant = default_max_relevant,
    default_theory_relevant = default_theory_relevant,
    explicit_forall = true,
    use_conjecture_for_hypotheses = use_conjecture_for_hypotheses}
 
 fun remotify_config system_name system_versions
-        ({proof_delims, known_failures, default_max_relevant_per_iter,
+        ({proof_delims, known_failures, default_max_relevant,
           default_theory_relevant, use_conjecture_for_hypotheses, ...}
          : prover_config) : prover_config =
   remote_config system_name system_versions proof_delims known_failures
-                default_max_relevant_per_iter default_theory_relevant
+                default_max_relevant default_theory_relevant
                 use_conjecture_for_hypotheses
 
 val remotify_name = prefix "remote_"
 fun remote_prover name system_name system_versions proof_delims known_failures
-                  default_max_relevant_per_iter default_theory_relevant
+                  default_max_relevant default_theory_relevant
                   use_conjecture_for_hypotheses =
   (remotify_name name,
    remote_config system_name system_versions proof_delims known_failures
-                 default_max_relevant_per_iter default_theory_relevant
+                 default_max_relevant default_theory_relevant
                  use_conjecture_for_hypotheses)
 fun remotify_prover (name, config) system_name system_versions =
   (remotify_name name, remotify_config system_name system_versions config)
 
 val remote_e = remotify_prover e "EP" ["1.0", "1.1", "1.2"]
-val remote_vampire = remotify_prover vampire "Vampire" ["9.9", "0.6", "1.0"]
+val remote_vampire = remotify_prover vampire "Vampire" ["9.0", "1.0", "0.6"]
 val remote_sine_e =
-  remote_prover "sine_e" "SInE" [] []
-                [(Unprovable, "says Unknown")] 150 (* FUDGE *) false true
+  remote_prover "sine_e" "SInE" [] [] [(Unprovable, "says Unknown")]
+                1000 (* FUDGE *) false true
 val remote_snark =
   remote_prover "snark" "SNARK---" [] [("refutation.", "end_refutation.")] []
-                50 (* FUDGE *) false true
+                350 (* FUDGE *) false true
 
 (* Setup *)
 
--- a/src/HOL/Tools/Datatype/datatype_abs_proofs.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_abs_proofs.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -416,7 +416,7 @@
     fun prove_case_cong ((t, nchotomy), case_rewrites) =
       let
         val (Const ("==>", _) $ tm $ _) = t;
-        val (Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, _) $ _ $ Ma)) = tm;
+        val (Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ Ma)) = tm;
         val cert = cterm_of thy;
         val nchotomy' = nchotomy RS spec;
         val [v] = Term.add_vars (concl_of nchotomy') [];
--- a/src/HOL/Tools/Datatype/datatype_aux.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_aux.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -120,8 +120,8 @@
 fun split_conj_thm th =
   ((th RS conjunct1)::(split_conj_thm (th RS conjunct2))) handle THM _ => [th];
 
-val mk_conj = foldr1 (HOLogic.mk_binop @{const_name "op &"});
-val mk_disj = foldr1 (HOLogic.mk_binop @{const_name "op |"});
+val mk_conj = foldr1 (HOLogic.mk_binop @{const_name HOL.conj});
+val mk_disj = foldr1 (HOLogic.mk_binop @{const_name HOL.disj});
 
 fun app_bnds t i = list_comb (t, map Bound (i - 1 downto 0));
 
--- a/src/HOL/Tools/Datatype/datatype_codegen.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_codegen.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -68,7 +68,7 @@
     val { descr, index, inject = inject_thms, distinct = distinct_thms, ... } =
       Datatype_Data.the_info thy tyco;
     val ty = Type (tyco, map TFree vs);
-    fun mk_eq (t1, t2) = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
+    fun mk_eq (t1, t2) = Const (@{const_name HOL.equal}, ty --> ty --> HOLogic.boolT)
       $ t1 $ t2;
     fun true_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.true_const);
     fun false_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.false_const);
@@ -83,7 +83,7 @@
     val distincts = maps prep_distinct (snd (nth (Datatype_Prop.make_distincts [descr] vs) index));
     val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty)));
     val simpset = Simplifier.global_context thy (HOL_basic_ss addsimps 
-      (map Simpdata.mk_eq (@{thm eq} :: @{thm eq_True} :: inject_thms @ distinct_thms)));
+      (map Simpdata.mk_eq (@{thm equal} :: @{thm eq_True} :: inject_thms @ distinct_thms)));
     fun prove prop = Skip_Proof.prove_global thy [] [] prop (K (ALLGOALS (simp_tac simpset)))
       |> Simpdata.mk_eq;
   in (map prove (triv_injects @ injects @ distincts), prove refl) end;
@@ -96,7 +96,7 @@
         fun mk_side const_name = Const (const_name, ty --> ty --> HOLogic.boolT)
           $ Free ("x", ty) $ Free ("y", ty);
         val def = HOLogic.mk_Trueprop (HOLogic.mk_eq
-          (mk_side @{const_name eq_class.eq}, mk_side @{const_name "op ="}));
+          (mk_side @{const_name HOL.equal}, mk_side @{const_name HOL.eq}));
         val def' = Syntax.check_term lthy def;
         val ((_, (_, thm)), lthy') = Specification.definition
           (NONE, (Attrib.empty_binding, def')) lthy;
@@ -115,7 +115,7 @@
       #> snd
   in
     thy
-    |> Class.instantiation (tycos, vs, [HOLogic.class_eq])
+    |> Class.instantiation (tycos, vs, [HOLogic.class_equal])
     |> fold_map add_def tycos
     |-> (fn def_thms => Class.prove_instantiation_exit_result (map o Morphism.thm)
          (fn _ => fn def_thms => tac def_thms) def_thms)
@@ -135,7 +135,7 @@
     val case_rewrites = maps (#case_rewrites o Datatype_Data.the_info thy) tycos;
     val certs = map (mk_case_cert thy) tycos;
     val tycos_eq = filter_out
-      (fn tyco => can (Sorts.mg_domain (Sign.classes_of thy) tyco) [HOLogic.class_eq]) tycos;
+      (fn tyco => can (Sorts.mg_domain (Sign.classes_of thy) tyco) [HOLogic.class_equal]) tycos;
   in
     if null css then thy
     else thy
--- a/src/HOL/Tools/Datatype/datatype_data.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_data.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -257,7 +257,9 @@
            Pretty.str " =" :: Pretty.brk 1 ::
            flat (separate [Pretty.brk 1, Pretty.str "| "]
              (map (single o pretty_constr) cos)));
-    in Thy_Output.output (Thy_Output.maybe_pretty_source (K pretty_datatype) src [()]) end);
+    in
+      Thy_Output.output ctxt (Thy_Output.maybe_pretty_source (K (K pretty_datatype)) ctxt src [()])
+    end);
 
 
 
@@ -428,7 +430,7 @@
           unflat rules (map Drule.zero_var_indexes_list raw_thms);
             (*FIXME somehow dubious*)
       in
-        ProofContext.theory_result
+        ProofContext.background_theory_result
           (prove_rep_datatype config dt_names alt_names descr vs
             raw_inject half_distinct raw_induct)
         #-> after_qed
--- a/src/HOL/Tools/Datatype/datatype_prop.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_prop.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -70,7 +70,7 @@
           val frees' = map Free ((map ((op ^) o (rpair "'")) tnames) ~~ Ts);
         in cons (HOLogic.mk_Trueprop (HOLogic.mk_eq
           (HOLogic.mk_eq (list_comb (constr_t, frees), list_comb (constr_t, frees')),
-           foldr1 (HOLogic.mk_binop @{const_name "op &"})
+           foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
              (map HOLogic.mk_eq (frees ~~ frees')))))
         end;
   in
@@ -149,7 +149,7 @@
     val prems = maps (fn ((i, (_, _, constrs)), T) =>
       map (make_ind_prem i T) constrs) (descr' ~~ recTs);
     val tnames = make_tnames recTs;
-    val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name "op &"})
+    val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
       (map (fn (((i, _), T), tname) => make_pred i T $ Free (tname, T))
         (descr' ~~ recTs ~~ tnames)))
 
--- a/src/HOL/Tools/Datatype/datatype_realizer.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_realizer.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -99,7 +99,7 @@
         if member (op =) is i then SOME
           (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Free (tname, T))
         else NONE) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names ~~ tnames));
-    val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name "op &"})
+    val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
       (map (fn ((((i, _), T), U), tname) =>
         make_pred i U T (mk_proj i is r) (Free (tname, T)))
           (descr ~~ recTs ~~ rec_result_Ts ~~ tnames)));
--- a/src/HOL/Tools/Function/function_common.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Function/function_common.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -164,7 +164,7 @@
 structure Termination_Simps = Named_Thms
 (
   val name = "termination_simp"
-  val description = "Simplification rule for termination proofs"
+  val description = "simplification rules for termination proofs"
 )
 
 
@@ -175,7 +175,7 @@
   type T = Proof.context -> tactic
   val empty = (fn _ => error "Termination prover not configured")
   val extend = I
-  fun merge (a, b) = b  (* FIXME ? *)
+  fun merge (a, _) = a
 )
 
 val set_termination_prover = TerminationProver.put
--- a/src/HOL/Tools/Function/function_core.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Function/function_core.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -860,9 +860,9 @@
           rtac (instantiate' [] [SOME (cterm_of thy lhs_acc)] case_split) 1
           THEN (rtac (Thm.forall_elim_vars 0 psimp) THEN_ALL_NEW assume_tac) 1
           THEN (simp_default_tac (simpset_of ctxt) 1)
-          THEN (etac not_acc_down 1)
-          THEN ((etac R_cases)
-            THEN_ALL_NEW (simp_default_tac (simpset_of ctxt))) 1)
+          THEN TRY ((etac not_acc_down 1)
+            THEN ((etac R_cases)
+              THEN_ALL_NEW (simp_default_tac (simpset_of ctxt))) 1))
         |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
       end
   in
--- a/src/HOL/Tools/Function/measure_functions.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Function/measure_functions.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -20,7 +20,7 @@
 (
   val name = "measure_function"
   val description =
-    "Rules that guide the heuristic generation of measure functions"
+    "rules that guide the heuristic generation of measure functions"
 );
 
 fun mk_is_measure t =
--- a/src/HOL/Tools/Function/scnp_reconstruct.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Function/scnp_reconstruct.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -68,7 +68,7 @@
   type T = multiset_setup option
   val empty = NONE
   val extend = I;
-  fun merge (v1, v2) = if is_some v2 then v2 else v1   (* FIXME prefer v1 !?! *)
+  fun merge (v1, v2) = if is_some v1 then v1 else v2
 )
 
 val multiset_setup = Multiset_Setup.put o SOME
--- a/src/HOL/Tools/Function/termination.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Function/termination.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -148,7 +148,7 @@
     val cs = Function_Lib.dest_binop_list @{const_name Lattices.sup} rel
     fun collect_pats (Const (@{const_name Collect}, _) $ Abs (_, _, c)) =
       let
-        val (Const (@{const_name "op &"}, _) $ (Const (@{const_name "op ="}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ _)
+        val (Const (@{const_name HOL.conj}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ _)
           = Term.strip_qnt_body @{const_name Ex} c
       in cons r o cons l end
   in
@@ -185,7 +185,7 @@
     val vs = Term.strip_qnt_vars @{const_name Ex} c
 
     (* FIXME: throw error "dest_call" for malformed terms *)
-    val (Const (@{const_name "op &"}, _) $ (Const (@{const_name "op ="}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ Gam)
+    val (Const (@{const_name HOL.conj}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ Gam)
       = Term.strip_qnt_body @{const_name Ex} c
     val (p, l') = dest_inj sk l
     val (q, r') = dest_inj sk r
--- a/src/HOL/Tools/Nitpick/minipick.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Nitpick/minipick.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -123,16 +123,16 @@
          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)
-       | Const (@{const_name "op ="}, _) $ t1 $ t2 =>
+       | Const (@{const_name HOL.eq}, _) $ t1 $ t2 =>
          RelEq (to_R_rep Ts t1, to_R_rep Ts t2)
        | Const (@{const_name ord_class.less_eq},
                 Type (@{type_name fun},
                       [Type (@{type_name fun}, [_, @{typ bool}]), _]))
          $ t1 $ t2 =>
          Subset (to_R_rep Ts t1, to_R_rep Ts t2)
-       | @{const "op &"} $ t1 $ t2 => And (to_F Ts t1, to_F Ts t2)
-       | @{const "op |"} $ t1 $ t2 => Or (to_F Ts t1, to_F Ts t2)
-       | @{const "op -->"} $ t1 $ t2 => Implies (to_F Ts t1, to_F Ts t2)
+       | @{const HOL.conj} $ t1 $ t2 => And (to_F Ts t1, to_F Ts t2)
+       | @{const HOL.disj} $ t1 $ t2 => Or (to_F Ts t1, to_F Ts t2)
+       | @{const HOL.implies} $ t1 $ t2 => Implies (to_F Ts t1, to_F Ts t2)
        | t1 $ t2 => Subset (to_S_rep Ts t2, to_R_rep Ts t1)
        | Free _ => raise SAME ()
        | Term.Var _ => raise SAME ()
@@ -165,20 +165,20 @@
          @{const Not} => to_R_rep Ts (eta_expand Ts t 1)
        | Const (@{const_name All}, _) => to_R_rep Ts (eta_expand Ts t 1)
        | Const (@{const_name Ex}, _) => to_R_rep Ts (eta_expand Ts t 1)
-       | Const (@{const_name "op ="}, _) $ _ => to_R_rep Ts (eta_expand Ts t 1)
-       | Const (@{const_name "op ="}, _) => to_R_rep Ts (eta_expand Ts t 2)
+       | Const (@{const_name HOL.eq}, _) $ _ => to_R_rep Ts (eta_expand Ts t 1)
+       | Const (@{const_name HOL.eq}, _) => to_R_rep Ts (eta_expand Ts t 2)
        | Const (@{const_name ord_class.less_eq},
                 Type (@{type_name fun},
                       [Type (@{type_name fun}, [_, @{typ bool}]), _])) $ _ =>
          to_R_rep Ts (eta_expand Ts t 1)
        | Const (@{const_name ord_class.less_eq}, _) =>
          to_R_rep Ts (eta_expand Ts t 2)
-       | @{const "op &"} $ _ => to_R_rep Ts (eta_expand Ts t 1)
-       | @{const "op &"} => to_R_rep Ts (eta_expand Ts t 2)
-       | @{const "op |"} $ _ => to_R_rep Ts (eta_expand Ts t 1)
-       | @{const "op |"} => to_R_rep Ts (eta_expand Ts t 2)
-       | @{const "op -->"} $ _ => to_R_rep Ts (eta_expand Ts t 1)
-       | @{const "op -->"} => to_R_rep Ts (eta_expand Ts t 2)
+       | @{const HOL.conj} $ _ => to_R_rep Ts (eta_expand Ts t 1)
+       | @{const HOL.conj} => to_R_rep Ts (eta_expand Ts t 2)
+       | @{const HOL.disj} $ _ => to_R_rep Ts (eta_expand Ts t 1)
+       | @{const HOL.disj} => to_R_rep Ts (eta_expand Ts t 2)
+       | @{const HOL.implies} $ _ => to_R_rep Ts (eta_expand Ts t 1)
+       | @{const HOL.implies} => to_R_rep Ts (eta_expand Ts t 2)
        | Const (@{const_name bot_class.bot},
                 T as Type (@{type_name fun}, [_, @{typ bool}])) =>
          empty_n_ary_rel (arity_of RRep card T)
--- a/src/HOL/Tools/Nitpick/nitpick.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -182,7 +182,7 @@
 fun none_true assigns = forall (not_equal (SOME true) o snd) assigns
 
 val syntactic_sorts =
-  @{sort "{default,zero,one,plus,minus,uminus,times,inverse,abs,sgn,ord,eq}"} @
+  @{sort "{default,zero,one,plus,minus,uminus,times,inverse,abs,sgn,ord,equal}"} @
   @{sort number}
 fun has_tfree_syntactic_sort (TFree (_, S as _ :: _)) =
     subset (op =) (S, syntactic_sorts)
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -386,13 +386,13 @@
     if t0 = conn_t then strip_connective t0 t2 @ strip_connective t0 t1 else [t]
   | strip_connective _ t = [t]
 fun strip_any_connective (t as (t0 $ _ $ _)) =
-    if t0 = @{const "op &"} orelse t0 = @{const "op |"} then
+    if t0 = @{const HOL.conj} orelse t0 = @{const HOL.disj} then
       (strip_connective t0 t, t0)
     else
       ([t], @{const Not})
   | strip_any_connective t = ([t], @{const Not})
-val conjuncts_of = strip_connective @{const "op &"}
-val disjuncts_of = strip_connective @{const "op |"}
+val conjuncts_of = strip_connective @{const HOL.conj}
+val disjuncts_of = strip_connective @{const HOL.disj}
 
 (* When you add constants to these lists, make sure to handle them in
    "Nitpick_Nut.nut_from_term", and perhaps in "Nitpick_Mono.consider_term" as
@@ -408,10 +408,10 @@
    (@{const_name True}, 0),
    (@{const_name All}, 1),
    (@{const_name Ex}, 1),
-   (@{const_name "op ="}, 1),
-   (@{const_name "op &"}, 2),
-   (@{const_name "op |"}, 2),
-   (@{const_name "op -->"}, 2),
+   (@{const_name HOL.eq}, 1),
+   (@{const_name HOL.conj}, 2),
+   (@{const_name HOL.disj}, 2),
+   (@{const_name HOL.implies}, 2),
    (@{const_name If}, 3),
    (@{const_name Let}, 2),
    (@{const_name Pair}, 2),
@@ -1275,7 +1275,7 @@
         forall is_Var args andalso not (has_duplicates (op =) args)
       | _ => false
     fun do_eq (Const (@{const_name "=="}, _) $ t1 $ _) = do_lhs t1
-      | do_eq (@{const Trueprop} $ (Const (@{const_name "op ="}, _) $ t1 $ _)) =
+      | do_eq (@{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ t1 $ _)) =
         do_lhs t1
       | do_eq _ = false
   in do_eq end
@@ -1347,7 +1347,7 @@
     @{const "==>"} $ _ $ t2 => term_under_def t2
   | Const (@{const_name "=="}, _) $ t1 $ _ => term_under_def t1
   | @{const Trueprop} $ t1 => term_under_def t1
-  | Const (@{const_name "op ="}, _) $ t1 $ _ => term_under_def t1
+  | Const (@{const_name HOL.eq}, _) $ t1 $ _ => term_under_def t1
   | Abs (_, _, t') => term_under_def t'
   | t1 $ _ => term_under_def t1
   | _ => t
@@ -1371,7 +1371,7 @@
     val (lhs, rhs) =
       case t of
         Const (@{const_name "=="}, _) $ t1 $ t2 => (t1, t2)
-      | @{const Trueprop} $ (Const (@{const_name "op ="}, _) $ t1 $ t2) =>
+      | @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ t1 $ t2) =>
         (t1, t2)
       | _ => raise TERM ("Nitpick_HOL.normalized_rhs_of", [t])
     val args = strip_comb lhs |> snd
@@ -1453,8 +1453,8 @@
   | @{const "==>"} $ _ $ t2 => lhs_of_equation t2
   | @{const Trueprop} $ t1 => lhs_of_equation t1
   | Const (@{const_name All}, _) $ Abs (_, _, t1) => lhs_of_equation t1
-  | Const (@{const_name "op ="}, _) $ t1 $ _ => SOME t1
-  | @{const "op -->"} $ _ $ t2 => lhs_of_equation t2
+  | Const (@{const_name HOL.eq}, _) $ t1 $ _ => SOME t1
+  | @{const HOL.implies} $ _ $ t2 => lhs_of_equation t2
   | _ => NONE
 fun is_constr_pattern _ (Bound _) = true
   | is_constr_pattern _ (Var _) = true
@@ -1807,7 +1807,7 @@
                         (betapply (t2, var_t))
     end
   | extensional_equal _ T t1 t2 =
-    Const (@{const_name "op ="}, T --> T --> bool_T) $ t1 $ t2
+    Const (@{const_name HOL.eq}, T --> T --> bool_T) $ t1 $ t2
 
 fun equationalize_term ctxt tag t =
   let
@@ -1816,7 +1816,7 @@
   in
     Logic.list_implies (prems,
         case concl of
-          @{const Trueprop} $ (Const (@{const_name "op ="}, Type (_, [T, _]))
+          @{const Trueprop} $ (Const (@{const_name HOL.eq}, Type (_, [T, _]))
                                $ t1 $ t2) =>
           @{const Trueprop} $ extensional_equal j T t1 t2
         | @{const Trueprop} $ t' =>
@@ -2148,8 +2148,8 @@
           fun repair_rec j (Const (@{const_name Ex}, T1) $ Abs (s2, T2, t2')) =
               Const (@{const_name Ex}, T1)
               $ Abs (s2, T2, repair_rec (j + 1) t2')
-            | repair_rec j (@{const "op &"} $ t1 $ t2) =
-              @{const "op &"} $ repair_rec j t1 $ repair_rec j t2
+            | repair_rec j (@{const HOL.conj} $ t1 $ t2) =
+              @{const HOL.conj} $ repair_rec j t1 $ repair_rec j t2
             | repair_rec j t =
               let val (head, args) = strip_comb t in
                 if head = Bound j then
@@ -2290,7 +2290,7 @@
   | simps => simps
 fun is_equational_fun_surely_complete hol_ctxt x =
   case equational_fun_axioms hol_ctxt x of
-    [@{const Trueprop} $ (Const (@{const_name "op ="}, _) $ t1 $ _)] =>
+    [@{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ t1 $ _)] =>
     strip_comb t1 |> snd |> forall is_Var
   | _ => false
 
--- a/src/HOL/Tools/Nitpick/nitpick_model.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_model.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -590,7 +590,7 @@
                       if co then
                         Const (@{const_name The}, (T --> bool_T) --> T)
                         $ Abs (cyclic_co_val_name (), T,
-                               Const (@{const_name "op ="}, T --> T --> bool_T)
+                               Const (@{const_name HOL.eq}, T --> T --> bool_T)
                                $ Bound 0 $ abstract_over (var, t))
                       else
                         cyclic_atom ()
@@ -849,7 +849,7 @@
 (** Model reconstruction **)
 
 fun unfold_outer_the_binders (t as Const (@{const_name The}, _)
-                                   $ Abs (s, T, Const (@{const_name "op ="}, _)
+                                   $ Abs (s, T, Const (@{const_name HOL.eq}, _)
                                                 $ Bound 0 $ t')) =
     betapply (Abs (s, T, t'), t) |> unfold_outer_the_binders
   | unfold_outer_the_binders t = t
--- a/src/HOL/Tools/Nitpick/nitpick_mono.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -222,7 +222,7 @@
   | fin_fun_body _ _ (t as Const (@{const_name None}, _)) = SOME t
   | fin_fun_body dom_T ran_T
                  ((t0 as Const (@{const_name If}, _))
-                  $ (t1 as Const (@{const_name "op ="}, _) $ Bound 0 $ t1')
+                  $ (t1 as Const (@{const_name HOL.eq}, _) $ Bound 0 $ t1')
                   $ t2 $ t3) =
     (if loose_bvar1 (t1', 0) then
        NONE
@@ -650,7 +650,7 @@
                                      Bound 0)))) accum
                   |>> mtype_of_mterm
                 end
-              | @{const_name "op ="} => do_equals T accum
+              | @{const_name HOL.eq} => do_equals T accum
               | @{const_name The} =>
                 (trace_msg (K "*** The"); raise UNSOLVABLE ())
               | @{const_name Eps} =>
@@ -760,7 +760,7 @@
                     do_term (incr_boundvars ~1 t1') accum
                   else
                     raise SAME ()
-                | (t11 as Const (@{const_name "op ="}, _)) $ Bound 0 $ t13 =>
+                | (t11 as Const (@{const_name HOL.eq}, _)) $ Bound 0 $ t13 =>
                   if not (loose_bvar1 (t13, 0)) then
                     do_term (incr_boundvars ~1 (t11 $ t13)) accum
                   else
@@ -774,10 +774,10 @@
                         (MAbs (s, T, M, S Minus, m'), accum |>> pop_bound)
                       end))
          | (t0 as Const (@{const_name All}, _))
-           $ Abs (s', T', (t10 as @{const "op -->"}) $ (t11 $ Bound 0) $ t12) =>
+           $ Abs (s', T', (t10 as @{const HOL.implies}) $ (t11 $ Bound 0) $ t12) =>
            do_bounded_quantifier t0 s' T' t10 t11 t12 accum
          | (t0 as Const (@{const_name Ex}, _))
-           $ Abs (s', T', (t10 as @{const "op &"}) $ (t11 $ Bound 0) $ t12) =>
+           $ Abs (s', T', (t10 as @{const HOL.conj}) $ (t11 $ Bound 0) $ t12) =>
            do_bounded_quantifier t0 s' T' t10 t11 t12 accum
          | Const (@{const_name Let}, _) $ t1 $ t2 =>
            do_term (betapply (t2, t1)) accum
@@ -876,19 +876,19 @@
                 do_term (@{const Not}
                          $ (HOLogic.eq_const (domain_type T0) $ t1
                             $ Abs (Name.uu, T1, @{const False}))) accum)
-           | Const (x as (@{const_name "op ="}, _)) $ t1 $ t2 =>
+           | Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 =>
              do_equals x t1 t2
            | Const (@{const_name Let}, _) $ t1 $ t2 =>
              do_formula sn (betapply (t2, t1)) accum
            | (t0 as Const (s0, _)) $ t1 $ t2 =>
              if s0 = @{const_name "==>"} orelse
                 s0 = @{const_name Pure.conjunction} orelse
-                s0 = @{const_name "op &"} orelse
-                s0 = @{const_name "op |"} orelse
-                s0 = @{const_name "op -->"} then
+                s0 = @{const_name HOL.conj} orelse
+                s0 = @{const_name HOL.disj} orelse
+                s0 = @{const_name HOL.implies} then
                let
                  val impl = (s0 = @{const_name "==>"} orelse
-                             s0 = @{const_name "op -->"})
+                             s0 = @{const_name HOL.implies})
                  val (m1, accum) = do_formula (sn |> impl ? negate) t1 accum
                  val (m2, accum) = do_formula sn t2 accum
                in
@@ -973,10 +973,10 @@
             do_conjunction t0 t1 t2 accum
           | (t0 as Const (@{const_name All}, _)) $ Abs (s0, T1, t1) =>
             do_all t0 s0 T1 t1 accum
-          | Const (x as (@{const_name "op ="}, _)) $ t1 $ t2 =>
+          | Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 =>
             consider_general_equals mdata true x t1 t2 accum
-          | (t0 as @{const "op &"}) $ t1 $ t2 => do_conjunction t0 t1 t2 accum
-          | (t0 as @{const "op -->"}) $ t1 $ t2 => do_implies t0 t1 t2 accum
+          | (t0 as @{const HOL.conj}) $ t1 $ t2 => do_conjunction t0 t1 t2 accum
+          | (t0 as @{const HOL.implies}) $ t1 $ t2 => do_implies t0 t1 t2 accum
           | _ => raise TERM ("Nitpick_Mono.consider_definitional_axiom.\
                              \do_formula", [t])
     in do_formula t end
@@ -1069,7 +1069,7 @@
                     Abs (Name.uu, set_T', @{const True})
                   | _ => Const (s, T')
                 else if s = @{const_name "=="} orelse
-                        s = @{const_name "op ="} then
+                        s = @{const_name HOL.eq} then
                   let
                     val T =
                       case T' of
--- a/src/HOL/Tools/Nitpick/nitpick_nut.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -447,7 +447,7 @@
                 val t1 = incr_boundvars n t1
                 val t2 = incr_boundvars n t2
                 val xs = map Bound (n - 1 downto 0)
-                val equation = Const (@{const_name "op ="},
+                val equation = Const (@{const_name HOL.eq},
                                       body_T --> body_T --> bool_T)
                                    $ betapplys (t1, xs) $ betapplys (t2, xs)
                 val t =
@@ -515,14 +515,14 @@
           do_description_operator The @{const_name undefined_fast_The} x t1
         | (Const (x as (@{const_name Eps}, _)), [t1]) =>
           do_description_operator Eps @{const_name undefined_fast_Eps} x t1
-        | (Const (@{const_name "op ="}, T), [t1]) =>
+        | (Const (@{const_name HOL.eq}, T), [t1]) =>
           Op1 (SingletonSet, range_type T, Any, sub t1)
-        | (Const (@{const_name "op ="}, T), [t1, t2]) => sub_equals T t1 t2
-        | (Const (@{const_name "op &"}, _), [t1, t2]) =>
+        | (Const (@{const_name HOL.eq}, T), [t1, t2]) => sub_equals T t1 t2
+        | (Const (@{const_name HOL.conj}, _), [t1, t2]) =>
           Op2 (And, bool_T, Any, sub' t1, sub' t2)
-        | (Const (@{const_name "op |"}, _), [t1, t2]) =>
+        | (Const (@{const_name HOL.disj}, _), [t1, t2]) =>
           Op2 (Or, bool_T, Any, sub t1, sub t2)
-        | (Const (@{const_name "op -->"}, _), [t1, t2]) =>
+        | (Const (@{const_name HOL.implies}, _), [t1, t2]) =>
           Op2 (Or, bool_T, Any, Op1 (Not, bool_T, Any, sub t1), sub' t2)
         | (Const (@{const_name If}, T), [t1, t2, t3]) =>
           Op3 (If, nth_range_type 3 T, Any, sub t1, sub t2, sub t3)
--- a/src/HOL/Tools/Nitpick/nitpick_preproc.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_preproc.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -41,9 +41,9 @@
     fun aux def (Const (@{const_name "=="}, _) $ t1 $ t2) =
         aux def t1 andalso aux false t2
       | aux def (@{const "==>"} $ t1 $ t2) = aux false t1 andalso aux def t2
-      | aux def (Const (@{const_name "op ="}, _) $ t1 $ t2) =
+      | aux def (Const (@{const_name HOL.eq}, _) $ t1 $ t2) =
         aux def t1 andalso aux false t2
-      | aux def (@{const "op -->"} $ t1 $ t2) = aux false t1 andalso aux def t2
+      | aux def (@{const HOL.implies} $ t1 $ t2) = aux false t1 andalso aux def t2
       | aux def (t1 $ t2) = aux def t1 andalso aux def t2
       | aux def (t as Const (s, _)) =
         (not def orelse t <> @{const Suc}) andalso
@@ -149,7 +149,7 @@
       case t of
         @{const Trueprop} $ t1 => box_var_in_def new_Ts old_Ts t1 z
       | Const (s0, _) $ t1 $ _ =>
-        if s0 = @{const_name "=="} orelse s0 = @{const_name "op ="} then
+        if s0 = @{const_name "=="} orelse s0 = @{const_name HOL.eq} then
           let
             val (t', args) = strip_comb t1
             val T' = fastype_of1 (new_Ts, do_term new_Ts old_Ts Neut t')
@@ -209,16 +209,16 @@
         do_quantifier new_Ts old_Ts polar s0 T0 s1 T1 t1
       | Const (s0 as @{const_name Ex}, T0) $ Abs (s1, T1, t1) =>
         do_quantifier new_Ts old_Ts polar s0 T0 s1 T1 t1
-      | Const (s0 as @{const_name "op ="}, T0) $ t1 $ t2 =>
+      | Const (s0 as @{const_name HOL.eq}, T0) $ t1 $ t2 =>
         do_equals new_Ts old_Ts s0 T0 t1 t2
-      | @{const "op &"} $ t1 $ t2 =>
-        @{const "op &"} $ do_term new_Ts old_Ts polar t1
+      | @{const HOL.conj} $ t1 $ t2 =>
+        @{const HOL.conj} $ do_term new_Ts old_Ts polar t1
         $ do_term new_Ts old_Ts polar t2
-      | @{const "op |"} $ t1 $ t2 =>
-        @{const "op |"} $ do_term new_Ts old_Ts polar t1
+      | @{const HOL.disj} $ t1 $ t2 =>
+        @{const HOL.disj} $ do_term new_Ts old_Ts polar t1
         $ do_term new_Ts old_Ts polar t2
-      | @{const "op -->"} $ t1 $ t2 =>
-        @{const "op -->"} $ do_term new_Ts old_Ts (flip_polarity polar) t1
+      | @{const HOL.implies} $ t1 $ t2 =>
+        @{const HOL.implies} $ do_term new_Ts old_Ts (flip_polarity polar) t1
         $ do_term new_Ts old_Ts polar t2
       | Const (x as (s, T)) =>
         if is_descr s then
@@ -332,9 +332,9 @@
         do_eq_or_imp Ts true def t0 t1 t2 seen
       | (t0 as @{const "==>"}) $ t1 $ t2 =>
         if def then (t, []) else do_eq_or_imp Ts false def t0 t1 t2 seen
-      | (t0 as Const (@{const_name "op ="}, _)) $ t1 $ t2 =>
+      | (t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2 =>
         do_eq_or_imp Ts true def t0 t1 t2 seen
-      | (t0 as @{const "op -->"}) $ t1 $ t2 =>
+      | (t0 as @{const HOL.implies}) $ t1 $ t2 =>
         do_eq_or_imp Ts false def t0 t1 t2 seen
       | Abs (s, T, t') =>
         let val (t', seen) = do_term (T :: Ts) def t' [] seen in
@@ -399,9 +399,9 @@
         aux_eq careful true t0 t1 t2
       | aux careful ((t0 as @{const "==>"}) $ t1 $ t2) =
         t0 $ aux false t1 $ aux careful t2
-      | aux careful ((t0 as Const (@{const_name "op ="}, _)) $ t1 $ t2) =
+      | aux careful ((t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2) =
         aux_eq careful true t0 t1 t2
-      | aux careful ((t0 as @{const "op -->"}) $ t1 $ t2) =
+      | aux careful ((t0 as @{const HOL.implies}) $ t1 $ t2) =
         t0 $ aux false t1 $ aux careful t2
       | aux careful (Abs (s, T, t')) = Abs (s, T, aux careful t')
       | aux careful (t1 $ t2) = aux careful t1 $ aux careful t2
@@ -449,7 +449,7 @@
 (** Destruction of universal and existential equalities **)
 
 fun curry_assms (@{const "==>"} $ (@{const Trueprop}
-                                   $ (@{const "op &"} $ t1 $ t2)) $ t3) =
+                                   $ (@{const HOL.conj} $ t1 $ t2)) $ t3) =
     curry_assms (Logic.list_implies ([t1, t2] |> map HOLogic.mk_Trueprop, t3))
   | curry_assms (@{const "==>"} $ t1 $ t2) =
     @{const "==>"} $ curry_assms t1 $ curry_assms t2
@@ -464,9 +464,9 @@
     and aux_implies prems zs t1 t2 =
       case t1 of
         Const (@{const_name "=="}, _) $ Var z $ t' => aux_eq prems zs z t' t1 t2
-      | @{const Trueprop} $ (Const (@{const_name "op ="}, _) $ Var z $ t') =>
+      | @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ Var z $ t') =>
         aux_eq prems zs z t' t1 t2
-      | @{const Trueprop} $ (Const (@{const_name "op ="}, _) $ t' $ Var z) =>
+      | @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ t' $ Var z) =>
         aux_eq prems zs z t' t1 t2
       | _ => aux (t1 :: prems) (Term.add_vars t1 zs) t2
     and aux_eq prems zs z t' t1 t2 =
@@ -499,7 +499,7 @@
             handle SAME () => do_term (t :: seen) ts
         in
           case t of
-            Const (@{const_name "op ="}, _) $ t1 $ t2 => do_eq true t1 t2
+            Const (@{const_name HOL.eq}, _) $ t1 $ t2 => do_eq true t1 t2
           | _ => do_term (t :: seen) ts
         end
   in do_term end
@@ -604,12 +604,12 @@
           do_quantifier s0 T0 s1 T1 t1
         | Const (s0 as @{const_name Ex}, T0) $ Abs (s1, T1, t1) =>
           do_quantifier s0 T0 s1 T1 t1
-        | @{const "op &"} $ t1 $ t2 =>
+        | @{const HOL.conj} $ t1 $ t2 =>
           s_conj (pairself (aux ss Ts js skolemizable polar) (t1, t2))
-        | @{const "op |"} $ t1 $ t2 =>
+        | @{const HOL.disj} $ t1 $ t2 =>
           s_disj (pairself (aux ss Ts js skolemizable polar) (t1, t2))
-        | @{const "op -->"} $ t1 $ t2 =>
-          @{const "op -->"} $ aux ss Ts js skolemizable (flip_polarity polar) t1
+        | @{const HOL.implies} $ t1 $ t2 =>
+          @{const HOL.implies} $ aux ss Ts js skolemizable (flip_polarity polar) t1
           $ aux ss Ts js skolemizable polar t2
         | (t0 as Const (@{const_name Let}, _)) $ t1 $ t2 =>
           t0 $ t1 $ aux ss Ts js skolemizable polar t2
@@ -620,8 +620,8 @@
             let
               val gfp = (fixpoint_kind_of_const thy def_table x = Gfp)
               val (pref, connective) =
-                if gfp then (lbfp_prefix, @{const "op |"})
-                else (ubfp_prefix, @{const "op &"})
+                if gfp then (lbfp_prefix, @{const HOL.disj})
+                else (ubfp_prefix, @{const HOL.conj})
               fun pos () = unrolled_inductive_pred_const hol_ctxt gfp x
                            |> aux ss Ts js skolemizable polar
               fun neg () = Const (pref ^ s, T)
@@ -653,7 +653,7 @@
 
 fun params_in_equation (@{const "==>"} $ _ $ t2) = params_in_equation t2
   | params_in_equation (@{const Trueprop} $ t1) = params_in_equation t1
-  | params_in_equation (Const (@{const_name "op ="}, _) $ t1 $ _) =
+  | params_in_equation (Const (@{const_name HOL.eq}, _) $ t1 $ _) =
     snd (strip_comb t1)
   | params_in_equation _ = []
 
@@ -1105,7 +1105,7 @@
   case t of
     (t0 as Const (@{const_name All}, T0)) $ Abs (s, T1, t1) =>
     (case t1 of
-       (t10 as @{const "op &"}) $ t11 $ t12 =>
+       (t10 as @{const HOL.conj}) $ t11 $ t12 =>
        t10 $ distribute_quantifiers (t0 $ Abs (s, T1, t11))
            $ distribute_quantifiers (t0 $ Abs (s, T1, t12))
      | (t10 as @{const Not}) $ t11 =>
@@ -1118,10 +1118,10 @@
          t0 $ Abs (s, T1, distribute_quantifiers t1))
   | (t0 as Const (@{const_name Ex}, T0)) $ Abs (s, T1, t1) =>
     (case distribute_quantifiers t1 of
-       (t10 as @{const "op |"}) $ t11 $ t12 =>
+       (t10 as @{const HOL.disj}) $ t11 $ t12 =>
        t10 $ distribute_quantifiers (t0 $ Abs (s, T1, t11))
            $ distribute_quantifiers (t0 $ Abs (s, T1, t12))
-     | (t10 as @{const "op -->"}) $ t11 $ t12 =>
+     | (t10 as @{const HOL.implies}) $ t11 $ t12 =>
        t10 $ distribute_quantifiers (Const (@{const_name All}, T0)
                                      $ Abs (s, T1, t11))
            $ distribute_quantifiers (t0 $ Abs (s, T1, t12))
--- a/src/HOL/Tools/Predicate_Compile/code_prolog.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/code_prolog.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -6,23 +6,39 @@
 
 signature CODE_PROLOG =
 sig
+  datatype prolog_system = SWI_PROLOG | YAP
+  type code_options =
+    {ensure_groundness : bool,
+     limited_types : (typ * int) list,
+     limited_predicates : (string list * int) list,
+     replacing : ((string * string) * string) list,
+     manual_reorder : ((string * int) * int list) list,
+     prolog_system : prolog_system}
+  val code_options_of : theory -> code_options 
+  val map_code_options : (code_options -> code_options) -> theory -> theory
+
   datatype arith_op = Plus | Minus
   datatype prol_term = Var of string | Cons of string | AppF of string * prol_term list
     | Number of int | ArithOp of arith_op * prol_term list;
   datatype prem = Conj of prem list
     | Rel of string * prol_term list | NotRel of string * prol_term list
     | Eq of prol_term * prol_term | NotEq of prol_term * prol_term
-    | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term;
-      
+    | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term
+    | Ground of string * typ;
+
   type clause = ((string * prol_term list) * prem);
   type logic_program = clause list;
   type constant_table = (string * string) list
-  
-  val generate : Proof.context -> string list -> (logic_program * constant_table)
+
+  val generate : bool -> Proof.context -> string -> (logic_program * constant_table)
   val write_program : logic_program -> string
-  val run : logic_program -> string -> string list -> int option -> prol_term list list
+  val run : prolog_system -> logic_program -> string -> string list -> int option -> prol_term list list
+
+  val quickcheck : Proof.context -> bool -> term -> int -> term list option * (bool list * bool)
 
   val trace : bool Unsynchronized.ref
+  
+  val replace : ((string * string) * string) -> logic_program -> logic_program
 end;
 
 structure Code_Prolog : CODE_PROLOG =
@@ -33,6 +49,45 @@
 val trace = Unsynchronized.ref false
 
 fun tracing s = if !trace then Output.tracing s else () 
+
+(* code generation options *)
+
+datatype prolog_system = SWI_PROLOG | YAP
+
+type code_options =
+  {ensure_groundness : bool,
+   limited_types : (typ * int) list,
+   limited_predicates : (string list * int) list,
+   replacing : ((string * string) * string) list,
+   manual_reorder : ((string * int) * int list) list,
+   prolog_system : prolog_system}
+
+structure Options = Theory_Data
+(
+  type T = code_options
+  val empty = {ensure_groundness = false,
+    limited_types = [], limited_predicates = [], replacing = [], manual_reorder = [],
+    prolog_system = SWI_PROLOG}
+  val extend = I;
+  fun merge
+    ({ensure_groundness = ensure_groundness1, limited_types = limited_types1,
+      limited_predicates = limited_predicates1, replacing = replacing1,
+      manual_reorder = manual_reorder1, prolog_system = prolog_system1},
+     {ensure_groundness = ensure_groundness2, limited_types = limited_types2,
+      limited_predicates = limited_predicates2, replacing = replacing2,
+      manual_reorder = manual_reorder2, prolog_system = prolog_system2}) =
+    {ensure_groundness = ensure_groundness1 orelse ensure_groundness2,
+     limited_types = AList.merge (op =) (K true) (limited_types1, limited_types2),
+     limited_predicates = AList.merge (op =) (K true) (limited_predicates1, limited_predicates2),
+     manual_reorder = AList.merge (op =) (K true) (manual_reorder1, manual_reorder2),
+     replacing = Library.merge (op =) (replacing1, replacing2),
+     prolog_system = prolog_system1};
+);
+
+val code_options_of = Options.get
+
+val map_code_options = Options.map
+
 (* general string functions *)
 
 val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode;
@@ -45,6 +100,21 @@
 datatype prol_term = Var of string | Cons of string | AppF of string * prol_term list
   | Number of int | ArithOp of arith_op * prol_term list;
 
+fun dest_Var (Var v) = v
+
+fun add_vars (Var v) = insert (op =) v
+  | add_vars (ArithOp (_, ts)) = fold add_vars ts
+  | add_vars (AppF (_, ts)) = fold add_vars ts
+  | add_vars _ = I
+
+fun map_vars f (Var v) = Var (f v)
+  | map_vars f (ArithOp (opr, ts)) = ArithOp (opr, map (map_vars f) ts)
+  | map_vars f (AppF (fs, ts)) = AppF (fs, map (map_vars f) ts)
+  | map_vars f t = t
+  
+fun maybe_AppF (c, []) = Cons c
+  | maybe_AppF (c, xs) = AppF (c, xs)
+
 fun is_Var (Var _) = true
   | is_Var _ = false
 
@@ -61,26 +131,61 @@
 datatype prem = Conj of prem list
   | Rel of string * prol_term list | NotRel of string * prol_term list
   | Eq of prol_term * prol_term | NotEq of prol_term * prol_term
-  | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term;
-    
+  | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term
+  | Ground of string * typ;
+
 fun dest_Rel (Rel (c, ts)) = (c, ts)
- 
+
+fun map_term_prem f (Conj prems) = Conj (map (map_term_prem f) prems)
+  | map_term_prem f (Rel (r, ts)) = Rel (r, map f ts)
+  | map_term_prem f (NotRel (r, ts)) = NotRel (r, map f ts)
+  | map_term_prem f (Eq (l, r)) = Eq (f l, f r)
+  | map_term_prem f (NotEq (l, r)) = NotEq (f l, f r)
+  | map_term_prem f (ArithEq (l, r)) = ArithEq (f l, f r)
+  | map_term_prem f (NotArithEq (l, r)) = NotArithEq (f l, f r)
+  | map_term_prem f (Ground (v, T)) = Ground (dest_Var (f (Var v)), T)
+
+fun fold_prem_terms f (Conj prems) = fold (fold_prem_terms f) prems
+  | fold_prem_terms f (Rel (_, ts)) = fold f ts
+  | fold_prem_terms f (NotRel (_, ts)) = fold f ts
+  | fold_prem_terms f (Eq (l, r)) = f l #> f r
+  | fold_prem_terms f (NotEq (l, r)) = f l #> f r
+  | fold_prem_terms f (ArithEq (l, r)) = f l #> f r
+  | fold_prem_terms f (NotArithEq (l, r)) = f l #> f r
+  | fold_prem_terms f (Ground (v, T)) = f (Var v)
+  
 type clause = ((string * prol_term list) * prem);
 
 type logic_program = clause list;
 
 (* translation from introduction rules to internal representation *)
 
+fun mk_conform f empty avoid name =
+  let
+    fun dest_Char (Symbol.Char c) = c
+    val name' = space_implode "" (map (dest_Char o Symbol.decode)
+      (filter (fn s => Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s)
+        (Symbol.explode name)))
+    val name'' = f (if name' = "" then empty else name')
+  in (if member (op =) avoid name'' then Name.variant avoid name'' else name'') end
+
 (** constant table **)
 
 type constant_table = (string * string) list
 
 (* assuming no clashing *)
-fun mk_constant_table consts =
-  AList.make (first_lower o Long_Name.base_name) consts
-
 fun declare_consts consts constant_table =
-  fold (fn c => AList.update (op =) (c, first_lower (Long_Name.base_name c))) consts constant_table
+  let
+    fun update' c table =
+      if AList.defined (op =) table c then table else
+        let
+          val c' = mk_conform first_lower "pred" (map snd table) (Long_Name.base_name c)
+        in
+          AList.update (op =) (c, c') table
+        end
+  in
+    fold update' consts constant_table
+  end
   
 fun translate_const constant_table c =
   case AList.lookup (op =) constant_table c of
@@ -96,16 +201,23 @@
   case inv_lookup (op =) constant_table c of
     SOME c' => c'
   | NONE => error ("No constant corresponding to "  ^ c)
-  
+
 (** translation of terms, literals, premises, and clauses **)
 
 fun translate_arith_const @{const_name "Groups.plus_class.plus"} = SOME Plus
   | translate_arith_const @{const_name "Groups.minus_class.minus"} = SOME Minus
   | translate_arith_const _ = NONE
 
+fun mk_nat_term constant_table n =
+  let
+    val zero = translate_const constant_table @{const_name "Groups.zero_class.zero"}
+    val Suc = translate_const constant_table @{const_name "Suc"}
+  in funpow n (fn t => AppF (Suc, [t])) (Cons zero) end
+
 fun translate_term ctxt constant_table t =
   case try HOLogic.dest_number t of
     SOME (@{typ "int"}, n) => Number n
+  | SOME (@{typ "nat"}, n) => mk_nat_term constant_table n
   | NONE =>
       (case strip_comb t of
         (Free (v, T), []) => Var v 
@@ -119,12 +231,12 @@
 
 fun translate_literal ctxt constant_table t =
   case strip_comb t of
-    (Const (@{const_name "op ="}, _), [l, r]) =>
+    (Const (@{const_name HOL.eq}, _), [l, r]) =>
       let
         val l' = translate_term ctxt constant_table l
         val r' = translate_term ctxt constant_table r
       in
-        (if is_Var l' andalso is_arith_term r' then ArithEq else Eq) (l', r')
+        (if is_Var l' andalso is_arith_term r' andalso not (is_Var r') then ArithEq else Eq) (l', r')
       end
   | (Const (c, _), args) =>
       Rel (translate_const constant_table c, map (translate_term ctxt constant_table) args)
@@ -134,11 +246,16 @@
   | NegRel_of (Eq eq) = NotEq eq
   | NegRel_of (ArithEq eq) = NotArithEq eq
 
-fun translate_prem ctxt constant_table t =  
+fun mk_groundness_prems t = map Ground (Term.add_frees t [])
+  
+fun translate_prem ensure_groundness ctxt constant_table t =  
     case try HOLogic.dest_not t of
-      SOME t => NegRel_of (translate_literal ctxt constant_table t)
+      SOME t =>
+        if ensure_groundness then
+          Conj (mk_groundness_prems t @ [NegRel_of (translate_literal ctxt constant_table t)])
+        else
+          NegRel_of (translate_literal ctxt constant_table t)
     | NONE => translate_literal ctxt constant_table t
-
     
 fun imp_prems_conv cv ct =
   case Thm.term_of ct of
@@ -156,64 +273,230 @@
       (Trueprop_conv (Conv.try_conv (Conv.rewr_conv @{thm Predicate.eq_is_eq}))))
     (Thm.transfer thy rule)
 
-fun translate_intros ctxt gr const constant_table =
+fun translate_intros ensure_groundness ctxt gr const constant_table =
   let
     val intros = map (preprocess_intro (ProofContext.theory_of ctxt)) (Graph.get_node gr const)
     val (intros', ctxt') = Variable.import_terms true (map prop_of intros) ctxt
     val constant_table' = declare_consts (fold Term.add_const_names intros' []) constant_table
+      |> declare_consts [@{const_name "Groups.zero_class.zero"}, @{const_name "Suc"}]
     fun translate_intro intro =
       let
         val head = HOLogic.dest_Trueprop (Logic.strip_imp_concl intro)
-        val prems = map HOLogic.dest_Trueprop (Logic.strip_imp_prems intro) 
-        val prems' = Conj (map (translate_prem ctxt' constant_table') prems)
+        val prems = map HOLogic.dest_Trueprop (Logic.strip_imp_prems intro)
+        val prems' = Conj (map (translate_prem ensure_groundness ctxt' constant_table') prems)
         val clause = (dest_Rel (translate_literal ctxt' constant_table' head), prems')
       in clause end
-  in (map translate_intro intros', constant_table') end
+    val res = (map translate_intro intros', constant_table')
+  in res end
+
+fun depending_preds_of (key, intros) =
+  fold Term.add_const_names (map Thm.prop_of intros) []
 
-fun generate ctxt const =
-  let 
-     fun strong_conn_of gr keys =
-      Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
-    val gr = Predicate_Compile_Core.intros_graph_of ctxt
-    val scc = strong_conn_of gr const
-    val constant_table = mk_constant_table (flat scc)
+fun add_edges edges_of key G =
+  let
+    fun extend' key (G, visited) = 
+      case try (Graph.get_node G) key of
+          SOME v =>
+            let
+              val new_edges = filter (fn k => is_some (try (Graph.get_node G) k)) (edges_of (key, v))
+              val (G', visited') = fold extend'
+                (subtract (op =) (key :: visited) new_edges) (G, key :: visited)
+            in
+              (fold (Graph.add_edge o (pair key)) new_edges G', visited')
+            end
+        | NONE => (G, visited)
   in
-    apfst flat (fold_map (translate_intros ctxt gr) (flat scc) constant_table)
+    fst (extend' key (G, []))
   end
 
-(* transform logic program *)
-
-(** ensure groundness of terms before negation **)
+fun generate ensure_groundness ctxt const =
+  let 
+    fun strong_conn_of gr keys =
+      Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
+    val gr = Predicate_Compile_Core.intros_graph_of ctxt
+    val gr' = add_edges depending_preds_of const gr
+    val scc = strong_conn_of gr' [const]
+    val constant_table = declare_consts (flat scc) []
+  in
+    apfst flat (fold_map (translate_intros ensure_groundness ctxt gr) (flat scc) constant_table)
+  end
+  
+(* implementation for fully enumerating predicates and
+  for size-limited predicates for enumerating the values of a datatype upto a specific size *)
 
-fun add_vars (Var x) vs = insert (op =) x vs
-  | add_vars (Cons c) vs = vs
-  | add_vars (AppF (f, args)) vs = fold add_vars args vs
+fun add_ground_typ (Conj prems) = fold add_ground_typ prems
+  | add_ground_typ (Ground (_, T)) = insert (op =) T
+  | add_ground_typ _ = I
+
+fun mk_relname (Type (Tcon, Targs)) =
+  first_lower (Long_Name.base_name Tcon) ^ space_implode "_" (map mk_relname Targs)
+  | mk_relname _ = raise Fail "unexpected type"
+
+fun mk_lim_relname T = "lim_" ^  mk_relname T
+
+(* This is copied from "pat_completeness.ML" *)
+fun inst_constrs_of thy (T as Type (name, _)) =
+  map (fn (Cn,CT) =>
+    Envir.subst_term_types (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT)))
+    (the (Datatype.get_constrs thy name))
+  | inst_constrs_of thy T = raise TYPE ("inst_constrs_of", [T], [])
 
-fun string_of_typ (Type (s, Ts)) = Long_Name.base_name s
+fun is_recursive_constr T (Const (constr_name, T')) = member (op =) (binder_types T') T
+  
+fun mk_ground_impl ctxt limited_types (T as Type (Tcon, Targs)) (seen, constant_table) =
+  if member (op =) seen T then ([], (seen, constant_table))
+  else
+    let
+      val (limited, size) = case AList.lookup (op =) limited_types T of
+        SOME s => (true, s)
+      | NONE => (false, 0)      
+      val rel_name = (if limited then mk_lim_relname else mk_relname) T
+      fun mk_impl (Const (constr_name, cT), recursive) (seen, constant_table) =
+        let
+          val constant_table' = declare_consts [constr_name] constant_table
+          val Ts = binder_types cT
+          val (rec_clauses, (seen', constant_table'')) =
+            fold_map (mk_ground_impl ctxt limited_types) Ts (seen, constant_table')
+          val vars = map (fn i => Var ("x" ^ string_of_int i)) (1 upto (length Ts))
+          val lim_var =
+            if limited then
+              if recursive then [AppF ("suc", [Var "Lim"])]              
+              else [Var "Lim"]
+            else [] 
+          fun mk_prem v T' =
+            if limited andalso T' = T then Rel (mk_lim_relname T', [Var "Lim", v])
+            else Rel (mk_relname T', [v])
+          val clause =
+            ((rel_name, lim_var @ [maybe_AppF (translate_const constant_table'' constr_name, vars)]),
+             Conj (map2 mk_prem vars Ts))
+        in
+          (clause :: flat rec_clauses, (seen', constant_table''))
+        end
+      val constrs = inst_constrs_of (ProofContext.theory_of ctxt) T
+      val constrs' = (constrs ~~ map (is_recursive_constr T) constrs)
+        |> (fn cs => filter_out snd cs @ filter snd cs)
+      val (clauses, constant_table') =
+        apfst flat (fold_map mk_impl constrs' (T :: seen, constant_table))
+      val size_term = funpow size (fn t => AppF ("suc", [t])) (Cons "zero")
+    in
+      ((if limited then
+        cons ((mk_relname T, [Var "x"]), Rel (mk_lim_relname T, [size_term, Var "x"]))
+      else I) clauses, constant_table')
+    end
+ | mk_ground_impl ctxt _ T (seen, constant_table) =
+   raise Fail ("unexpected type :" ^ Syntax.string_of_typ ctxt T)
 
-fun mk_groundness_prems ts =
+fun replace_ground (Conj prems) = Conj (map replace_ground prems)
+  | replace_ground (Ground (x, T)) =
+    Rel (mk_relname T, [Var x])  
+  | replace_ground p = p
+  
+fun add_ground_predicates ctxt limited_types (p, constant_table) =
   let
-    val vars = fold add_vars ts []
-    fun mk_ground v =
-      Rel ("ground", [Var v])
+    val ground_typs = fold (add_ground_typ o snd) p []
+    val (grs, (_, constant_table')) = fold_map (mk_ground_impl ctxt limited_types) ground_typs ([], constant_table)
+    val p' = map (apsnd replace_ground) p
   in
-    map mk_ground vars
+    ((flat grs) @ p', constant_table')
   end
 
-fun ensure_groundness_prem (NotRel (c, ts)) = Conj (mk_groundness_prems ts @ [NotRel (c, ts)]) 
-  | ensure_groundness_prem (NotEq (l, r)) = Conj (mk_groundness_prems [l, r] @ [NotEq (l, r)])
-  | ensure_groundness_prem (Conj ps) = Conj (map ensure_groundness_prem ps)
-  | ensure_groundness_prem p = p
+(* make depth-limited version of predicate *)
+
+fun mk_lim_rel_name rel_name = "lim_" ^ rel_name
+
+fun mk_depth_limited rel_names ((rel_name, ts), prem) =
+  let
+    fun has_positive_recursive_prems (Conj prems) = exists has_positive_recursive_prems prems
+      | has_positive_recursive_prems (Rel (rel, ts)) = member (op =) rel_names rel
+      | has_positive_recursive_prems _ = false
+    fun mk_lim_prem (Conj prems) = Conj (map mk_lim_prem prems)
+      | mk_lim_prem (p as Rel (rel, ts)) =
+        if member (op =) rel_names rel then Rel (mk_lim_rel_name rel, Var "Lim" :: ts) else p
+      | mk_lim_prem p = p
+  in
+    if has_positive_recursive_prems prem then
+      ((mk_lim_rel_name rel_name, (AppF ("suc", [Var "Lim"]))  :: ts), mk_lim_prem prem)
+    else
+      ((mk_lim_rel_name rel_name, (Var "Lim") :: ts), prem)
+  end
+
+fun add_limited_predicates limited_predicates =
+  let                                     
+    fun add (rel_names, limit) (p, constant_table) = 
+      let
+        val clauses = filter (fn ((rel, _), _) => member (op =) rel_names rel) p
+        val clauses' = map (mk_depth_limited rel_names) clauses
+        fun nat_term_of n = funpow n (fn t => AppF ("suc", [t])) (Cons "zero")
+        fun mk_entry_clause rel_name =
+          let
+            val nargs = length (snd (fst
+              (the (find_first (fn ((rel, _), _) => rel = rel_name) clauses))))
+            val vars = map (fn i => Var ("x" ^ string_of_int i)) (1 upto nargs)        
+          in
+            (("limited_" ^ rel_name, vars), Rel ("lim_" ^ rel_name, nat_term_of limit :: vars))
+          end
+      in (p @ (map mk_entry_clause rel_names) @ clauses', constant_table) end
+  in
+    fold add limited_predicates
+  end
+
+
+(* replace predicates in clauses *)
 
-fun ensure_groundness_before_negation p =
-  map (apsnd ensure_groundness_prem) p
+(* replace (A, B, C) p = replace A by B in clauses of C *)
+fun replace ((from, to), location) p =
+  let
+    fun replace_prem (Conj prems) = Conj (map replace_prem prems)
+      | replace_prem (r as Rel (rel, ts)) =
+          if rel = from then Rel (to, ts) else r
+      | replace_prem r = r
+  in
+    map (fn ((rel, args), prem) => ((rel, args), (if rel = location then replace_prem else I) prem)) p
+  end
+
+  
+(* reorder manually : reorder premises of ith clause of predicate p by a permutation perm *)
+
+fun reorder_manually reorder p =
+  let
+    fun reorder' (clause as ((rel, args), prem)) seen =
+      let
+        val seen' = AList.map_default (op =) (rel, 0) (fn x => x + 1) seen
+        val i = the (AList.lookup (op =) seen' rel)
+        val perm = AList.lookup (op =) reorder (rel, i)
+        val prem' = (case perm of 
+          SOME p => (case prem of Conj prems => Conj (map (nth prems) p) | _ => prem)
+        | NONE => prem)
+      in (((rel, args), prem'), seen') end
+  in
+    fst (fold_map reorder' p [])
+  end
+(* rename variables to prolog-friendly names *)
+
+fun rename_vars_term renaming = map_vars (fn v => the (AList.lookup (op =) renaming v))
+
+fun rename_vars_prem renaming = map_term_prem (rename_vars_term renaming)
+
+fun is_prolog_conform v =
+  forall (fn s => Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s) (Symbol.explode v)
+  
+fun mk_renaming v renaming =
+  (v, mk_conform first_upper "Var" (map snd renaming) v) :: renaming
+
+fun rename_vars_clause ((rel, args), prem) =
+  let
+    val vars = fold_prem_terms add_vars prem (fold add_vars args [])
+    val renaming = fold mk_renaming vars []
+  in ((rel, map (rename_vars_term renaming) args), rename_vars_prem renaming prem) end
+  
+val rename_vars_program = map rename_vars_clause
 
 (* code printer *)
 
 fun write_arith_op Plus = "+"
   | write_arith_op Minus = "-"
 
-fun write_term (Var v) = first_upper v
+fun write_term (Var v) = v
   | write_term (Cons c) = c
   | write_term (AppF (f, args)) = f ^ "(" ^ space_implode ", " (map write_term args) ^ ")"
   | write_term (ArithOp (oper, [a1, a2])) = write_term a1 ^ " " ^ write_arith_op oper ^ " " ^ write_term a2
@@ -236,14 +519,16 @@
 fun write_program p =
   cat_lines (map write_clause p) 
 
-(** query templates **)
+(* query templates *)
 
-fun query_first rel vnames =
+(** query and prelude for swi-prolog **)
+
+fun swi_prolog_query_first rel vnames =
   "eval :- once("  ^ rel ^ "(" ^ space_implode ", " vnames ^ ")),\n" ^
   "writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^
   "\\n', [" ^ space_implode ", " vnames ^ "]).\n"
   
-fun query_firstn n rel vnames =
+fun swi_prolog_query_firstn n rel vnames =
   "eval :- findnsols(" ^ string_of_int n ^ ", (" ^ space_implode ", " vnames ^ "), " ^
     rel ^ "(" ^ space_implode ", " vnames ^ "), Sols), writelist(Sols).\n" ^
     "writelist([]).\n" ^
@@ -251,19 +536,52 @@
     "writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^
     "\\n', [" ^ space_implode ", " vnames ^ "]), writelist(T).\n"
   
-val prelude =
+val swi_prolog_prelude =
   "#!/usr/bin/swipl -q -t main -f\n\n" ^
   ":- use_module(library('dialect/ciao/aggregates')).\n" ^
-  ":- style_check(-singleton).\n\n" ^
+  ":- style_check(-singleton).\n" ^
+  ":- style_check(-discontiguous).\n" ^ 	
+  ":- style_check(-atom).\n\n" ^
   "main :- catch(eval, E, (print_message(error, E), fail)), halt.\n" ^
   "main :- halt(1).\n"
 
+(** query and prelude for yap **)
+
+fun yap_query_first rel vnames =
+  "eval :- once(" ^ rel ^ "(" ^ space_implode ", " vnames ^ ")),\n" ^
+  "format('" ^ space_implode ";" (map (fn v => v ^ " = ~w") vnames) ^
+  "\\n', [" ^ space_implode ", " vnames ^ "]).\n"
+
+val yap_prelude =
+  "#!/usr/bin/yap -L\n\n" ^
+  ":- initialization(eval).\n"
+
+(* system-dependent query, prelude and invocation *)
+
+fun query system nsols = 
+  case system of
+    SWI_PROLOG =>
+      (case nsols of NONE => swi_prolog_query_first | SOME n => swi_prolog_query_firstn n)
+  | YAP =>
+      case nsols of NONE => yap_query_first | SOME n =>
+        error "No support for querying multiple solutions in the prolog system yap"
+
+fun prelude system =
+  case system of SWI_PROLOG => swi_prolog_prelude | YAP => yap_prelude
+
+fun invoke system file_name =
+  let
+    val cmd =
+      case system of SWI_PROLOG => "/usr/local/bin/swipl -f " | YAP => "/usr/local/bin/yap -L "
+  in fst (bash_output (cmd ^ file_name)) end
+
 (* parsing prolog solution *)
+
 val scan_number =
   Scan.many1 Symbol.is_ascii_digit
 
 val scan_atom =
-  Scan.many1 (fn s => Symbol.is_ascii_lower s orelse Symbol.is_ascii_quasi s)
+  Scan.many1 (fn s => Symbol.is_ascii_lower s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s)
 
 val scan_var =
   Scan.many1
@@ -303,28 +621,30 @@
         (l :: r :: []) => parse_term (unprefix " " r)
       | _ => raise Fail "unexpected equation in prolog output"
     fun parse_solution s = map dest_eq (space_explode ";" s)
+    val sols = case space_explode "\n" sol of [] => [] | s => fst (split_last s)  
   in
-    map parse_solution (fst (split_last (space_explode "\n" sol)))
+    map parse_solution sols
   end 
   
 (* calling external interpreter and getting results *)
 
-fun run p query_rel vnames nsols =
+fun run system p query_rel vnames nsols =
   let
-    val cmd = Path.named_root
-    val query = case nsols of NONE => query_first | SOME n => query_firstn n 
-    val prog = prelude ^ query query_rel vnames ^ write_program p
+    val p' = rename_vars_program p
+    val _ = tracing "Renaming variable names..."
+    val renaming = fold mk_renaming vnames [] 
+    val vnames' = map (fn v => the (AList.lookup (op =) renaming v)) vnames
+    val prog = prelude system ^ query system nsols query_rel vnames' ^ write_program p'
     val _ = tracing ("Generated prolog program:\n" ^ prog)
-    val prolog_file = File.tmp_path (Path.basic "prolog_file")
-    val _ = File.write prolog_file prog
-    val (solution, _) = bash_output ("/usr/local/bin/swipl -f " ^ File.shell_path prolog_file)
+    val solution = Cache_IO.with_tmp_file "prolog_file" (fn prolog_file =>
+      (File.write prolog_file prog; invoke system (Path.implode prolog_file)))
     val _ = tracing ("Prolog returned solution(s):\n" ^ solution)
     val tss = parse_solutions solution
   in
     tss
   end
 
-(* values command *)
+(* restoring types in terms *)
 
 fun restore_term ctxt constant_table (Var s, T) = Free (s, T)
   | restore_term ctxt constant_table (Number n, @{typ "int"}) = HOLogic.mk_number @{typ "int"} n
@@ -343,8 +663,33 @@
         map (restore_term ctxt constant_table) (args ~~ argsT'))
     end
 
+(* values command *)
+
+val preprocess_options = Predicate_Compile_Aux.Options {
+  expected_modes = NONE,
+  proposed_modes = NONE,
+  proposed_names = [],
+  show_steps = false,
+  show_intermediate_results = false,
+  show_proof_trace = false,
+  show_modes = false,
+  show_mode_inference = false,
+  show_compilation = false,
+  show_caught_failures = false,
+  skip_proof = true,
+  no_topmost_reordering = false,
+  function_flattening = true,
+  specialise = false,
+  fail_safe_function_flattening = false,
+  no_higher_order_predicate = [],
+  inductify = false,
+  detect_switches = true,
+  compilation = Predicate_Compile_Aux.Pred
+}
+
 fun values ctxt soln t_compr =
   let
+    val options = code_options_of (ProofContext.theory_of ctxt)
     val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
       | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term ctxt t_compr);
     val (body, Ts, fp) = HOLogic.strip_psplits split;
@@ -360,10 +705,24 @@
       case try (map (fst o dest_Free)) all_args of
         SOME vs => vs
       | NONE => error ("Not only free variables in " ^ commas (map (Syntax.string_of_term ctxt) all_args))
+    val _ = tracing "Preprocessing specification..."
+    val T = Sign.the_const_type (ProofContext.theory_of ctxt) name
+    val t = Const (name, T)
+    val thy' =
+      Theory.copy (ProofContext.theory_of ctxt)
+      |> Predicate_Compile.preprocess preprocess_options t
+    val ctxt' = ProofContext.init_global thy'
     val _ = tracing "Generating prolog program..."
-    val (p, constant_table) = generate ctxt [name]
+    val (p, constant_table) = generate (#ensure_groundness options) ctxt' name
+      |> (if #ensure_groundness options then
+          add_ground_predicates ctxt' (#limited_types options)
+        else I)
+      |> add_limited_predicates (#limited_predicates options)
+      |> apfst (fold replace (#replacing options))
+      |> apfst (reorder_manually (#manual_reorder options))
     val _ = tracing "Running prolog program..."
-    val tss = run p (translate_const constant_table name) (map first_upper vnames) soln
+    val tss = run (#prolog_system options)
+      p (translate_const constant_table name) (map first_upper vnames) soln
     val _ = tracing "Restoring terms..."
     val empty = Const("Orderings.bot_class.bot", fastype_of t_compr)
     fun mk_insert x S =
@@ -379,17 +738,18 @@
             mk_set_compr (t :: in_insert) ts xs
           else
             let
-              val uu as (uuN, uuT) = singleton (Variable.variant_frees ctxt [t]) ("uu", fastype_of t)
+              val uu as (uuN, uuT) = singleton (Variable.variant_frees ctxt' [t]) ("uu", fastype_of t)
               val set_compr =
                 HOLogic.mk_Collect (uuN, uuT, fold (fn (s, T) => fn t => HOLogic.mk_exists (s, T, t))
                   frees (HOLogic.mk_conj (HOLogic.mk_eq (Free uu, t), @{term "True"})))
             in
-              set_compr :: (if null in_insert then xs else (fold mk_insert in_insert empty) :: xs)  
+              mk_set_compr [] ts
+                (set_compr :: (if null in_insert then xs else (fold mk_insert in_insert empty) :: xs))  
             end
         end
   in
       foldl1 (HOLogic.mk_binop @{const_name sup}) (mk_set_compr []
-        (map (fn ts => HOLogic.mk_tuple (map (restore_term ctxt constant_table) (ts ~~ Ts))) tss) [])
+        (map (fn ts => HOLogic.mk_tuple (map (restore_term ctxt' constant_table) (ts ~~ Ts))) tss) [])
   end
 
 fun values_cmd print_modes soln raw_t state =
@@ -416,4 +776,57 @@
    >> (fn ((print_modes, soln), t) => Toplevel.keep
         (values_cmd print_modes soln t))); (*FIXME does not preserve the previous functionality*)
 
+(* quickcheck generator *)
+
+(* FIXME: large copy of Predicate_Compile_Quickcheck - refactor out commons *)
+
+fun strip_imp_prems (Const(@{const_name HOL.implies}, _) $ A $ B) = A :: strip_imp_prems B
+  | strip_imp_prems _ = [];
+
+fun strip_imp_concl (Const(@{const_name HOL.implies}, _) $ A $ B) = strip_imp_concl B
+  | strip_imp_concl A = A : term;
+
+fun strip_horn A = (strip_imp_prems A, strip_imp_concl A);
+
+fun quickcheck ctxt report t size =
+  let
+    val options = code_options_of (ProofContext.theory_of ctxt)
+    val thy = Theory.copy (ProofContext.theory_of ctxt)
+    val (vs, t') = strip_abs t
+    val vs' = Variable.variant_frees ctxt [] vs
+    val Ts = map snd vs'
+    val t'' = subst_bounds (map Free (rev vs'), t')
+    val (prems, concl) = strip_horn t''
+    val constname = "quickcheck"
+    val full_constname = Sign.full_bname thy constname
+    val constT = Ts ---> @{typ bool}
+    val thy1 = Sign.add_consts_i [(Binding.name constname, constT, NoSyn)] thy
+    val const = Const (full_constname, constT)
+    val t = Logic.list_implies
+      (map HOLogic.mk_Trueprop (prems @ [HOLogic.mk_not concl]),
+       HOLogic.mk_Trueprop (list_comb (Const (full_constname, constT), map Free vs')))
+    val tac = fn _ => Skip_Proof.cheat_tac thy1
+    val intro = Goal.prove (ProofContext.init_global thy1) (map fst vs') [] t tac
+    val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1
+    val thy3 = Predicate_Compile.preprocess preprocess_options const thy2
+    val ctxt' = ProofContext.init_global thy3
+    val _ = tracing "Generating prolog program..."
+    val (p, constant_table) = generate true ctxt' full_constname
+      |> add_ground_predicates ctxt' (#limited_types options)
+      |> add_limited_predicates (#limited_predicates options)
+      |> apfst (fold replace (#replacing options))
+      |> apfst (reorder_manually (#manual_reorder options))
+    val _ = tracing "Running prolog program..."
+    val tss = run (#prolog_system options)
+      p (translate_const constant_table full_constname) (map fst vs') (SOME 1)
+    val _ = tracing "Restoring terms..."
+    val res =
+      case tss of
+        [ts] => SOME (map (restore_term ctxt' constant_table) (ts ~~ Ts))
+      | _ => NONE
+    val empty_report = ([], false)
+  in
+    (res, empty_report)
+  end; 
+
 end;
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -176,9 +176,9 @@
      val t = Const (const, T)
      val options = extract_options (((expected_modes, proposed_modes), raw_options), const)
   in
-    if (is_inductify options) then
+    if is_inductify options then
       let
-        val lthy' = Local_Theory.theory (preprocess options t) lthy
+        val lthy' = Local_Theory.background_theory (preprocess options t) lthy
         val const =
           case Predicate_Compile_Fun.pred_of_function (ProofContext.theory_of lthy') const of
             SOME c => c
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -405,13 +405,13 @@
 (* general syntactic functions *)
 
 (*Like dest_conj, but flattens conjunctions however nested*)
-fun conjuncts_aux (Const (@{const_name "op &"}, _) $ t $ t') conjs = conjuncts_aux t (conjuncts_aux t' conjs)
+fun conjuncts_aux (Const (@{const_name HOL.conj}, _) $ t $ t') conjs = conjuncts_aux t (conjuncts_aux t' conjs)
   | conjuncts_aux t conjs = t::conjs;
 
 fun conjuncts t = conjuncts_aux t [];
 
 fun is_equationlike_term (Const ("==", _) $ _ $ _) = true
-  | is_equationlike_term (Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, _) $ _ $ _)) = true
+  | is_equationlike_term (Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ _)) = true
   | is_equationlike_term _ = false
   
 val is_equationlike = is_equationlike_term o prop_of 
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -63,6 +63,19 @@
   val add_random_dseq_equations : options -> string list -> theory -> theory
   val add_new_random_dseq_equations : options -> string list -> theory -> theory
   val mk_tracing : string -> term -> term
+  val prepare_intrs : options -> compilation -> theory -> string list -> thm list ->
+    ((string * typ) list * string list * string list * (string * mode list) list *
+      (string *  (Term.term list * Predicate_Compile_Aux.indprem list) list) list)
+  type mode_analysis_options = {use_random : bool, reorder_premises : bool, infer_pos_and_neg_modes : bool}  
+  datatype mode_derivation = Mode_App of mode_derivation * mode_derivation | Context of mode
+  | Mode_Pair of mode_derivation * mode_derivation | Term of mode
+  type moded_clause = term list * (Predicate_Compile_Aux.indprem * mode_derivation) list
+  type 'a pred_mode_table = (string * ((bool * mode) * 'a) list) list
+
+  val infer_modes : 
+    mode_analysis_options -> options -> compilation -> (string * typ) list -> (string * mode list) list ->
+      string list -> (string *  (Term.term list * Predicate_Compile_Aux.indprem list) list) list ->
+      theory -> ((moded_clause list pred_mode_table * string list) * theory)
 end;
 
 structure Predicate_Compile_Core : PREDICATE_COMPILE_CORE =
@@ -524,7 +537,7 @@
 
 fun dest_conjunct_prem th =
   case HOLogic.dest_Trueprop (prop_of th) of
-    (Const (@{const_name "op &"}, _) $ t $ t') =>
+    (Const (@{const_name HOL.conj}, _) $ t $ t') =>
       dest_conjunct_prem (th RS @{thm conjunct1})
         @ dest_conjunct_prem (th RS @{thm conjunct2})
     | _ => [th]
@@ -587,7 +600,7 @@
 
 fun preprocess_elim ctxt elimrule =
   let
-    fun replace_eqs (Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, T) $ lhs $ rhs)) =
+    fun replace_eqs (Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, T) $ lhs $ rhs)) =
        HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs)
      | replace_eqs t = t
     val thy = ProofContext.theory_of ctxt
@@ -730,9 +743,7 @@
   type T = (mode * (compilation_funs -> typ -> term)) list Symtab.table;
   val empty = Symtab.empty;
   val extend = I;
-  val merge = Symtab.merge ((K true)
-    : ((mode * (compilation_funs -> typ -> term)) list *
-      (mode * (compilation_funs -> typ -> term)) list -> bool));
+  fun merge data : T = Symtab.merge (K true) data;
 );
 
 fun alternative_compilation_of_global thy pred_name mode =
@@ -3033,12 +3044,13 @@
     "adding alternative introduction rules for code generation of inductive predicates"
 
 (* TODO: make Theory_Data to Generic_Data & remove duplication of local theory and theory *)
+(* FIXME ... this is important to avoid changing the background theory below *)
 fun generic_code_pred prep_const options raw_const lthy =
   let
     val thy = ProofContext.theory_of lthy
     val const = prep_const thy raw_const
     val ctxt = ProofContext.init_global thy
-    val lthy' = Local_Theory.theory (PredData.map
+    val lthy' = Local_Theory.background_theory (PredData.map
         (extend (fetch_pred_data ctxt) (depending_preds_of ctxt) const)) lthy
     val thy' = ProofContext.theory_of lthy'
     val ctxt' = ProofContext.init_global thy'
@@ -3063,7 +3075,7 @@
         val global_thms = ProofContext.export goal_ctxt
           (ProofContext.init_global (ProofContext.theory_of goal_ctxt)) (map the_single thms)
       in
-        goal_ctxt |> Local_Theory.theory (fold set_elim global_thms #>
+        goal_ctxt |> Local_Theory.background_theory (fold set_elim global_thms #>
           ((case compilation options of
              Pred => add_equations
            | DSeq => add_dseq_equations
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -111,7 +111,7 @@
 
 fun mk_meta_equation th =
   case prop_of th of
-    Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, _) $ _ $ _) => th RS @{thm eq_reflection}
+    Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ _) => th RS @{thm eq_reflection}
   | _ => th
 
 val meta_fun_cong = @{lemma "f == g ==> f x == g x" by simp}
@@ -217,12 +217,12 @@
    @{const_name "==>"},
    @{const_name Trueprop},
    @{const_name Not},
-   @{const_name "op ="},
-   @{const_name "op -->"},
+   @{const_name HOL.eq},
+   @{const_name HOL.implies},
    @{const_name All},
    @{const_name Ex}, 
-   @{const_name "op &"},
-   @{const_name "op |"}]
+   @{const_name HOL.conj},
+   @{const_name HOL.disj}]
 
 fun special_cases (c, T) = member (op =) [
   @{const_name Product_Type.Unity},
@@ -275,7 +275,6 @@
       else
         let
           val specs = get_specification options thy t
-            (*|> Predicate_Compile_Set.unfold_set_notation*)
           (*val _ = print_specification options thy constname specs*)
           val us = defiants_of specs
         in
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -21,8 +21,7 @@
 structure Fun_Pred = Theory_Data
 (
   type T = (term * term) Item_Net.T;
-  val empty = Item_Net.init ((op aconv o pairself fst) : (term * term) * (term * term) -> bool)
-    (single o fst);
+  val empty : T = Item_Net.init (op aconv o pairself fst) (single o fst);
   val extend = I;
   val merge = Item_Net.merge;
 )
@@ -352,13 +351,17 @@
         |> map (fn (resargs, (names', prems')) =>
           let
             val prem' = HOLogic.mk_Trueprop (mk_lit (list_comb (P, resargs)))
-          in (prem'::prems', names') end)
+          in (prems' @ [prem'], names') end)
       end
     val intro_ts' = folds_map rewrite prems frees
       |> maps (fn (prems', frees') =>
         rewrite concl frees'
-        |> map (fn (concl'::conclprems, _) =>
-          Logic.list_implies ((flat prems') @ conclprems, concl')))
+        |> map (fn (conclprems, _) =>
+          let
+            val (conclprems', concl') = split_last conclprems
+          in
+            Logic.list_implies ((flat prems') @ conclprems', concl')
+          end))
     (*val _ = tracing ("Rewritten intro to " ^
       commas (map (Syntax.string_of_term_global thy) intro_ts'))*)
   in
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -89,8 +89,8 @@
 fun is_compound ((Const (@{const_name Not}, _)) $ t) =
     error "is_compound: Negation should not occur; preprocessing is defect"
   | is_compound ((Const (@{const_name Ex}, _)) $ _) = true
-  | is_compound ((Const (@{const_name "op |"}, _)) $ _ $ _) = true
-  | is_compound ((Const (@{const_name "op &"}, _)) $ _ $ _) =
+  | is_compound ((Const (@{const_name HOL.disj}, _)) $ _ $ _) = true
+  | is_compound ((Const (@{const_name HOL.conj}, _)) $ _ $ _) =
     error "is_compound: Conjunction should not occur; preprocessing is defect"
   | is_compound _ = false
 
@@ -250,7 +250,7 @@
 
 fun split_conjs thy t =
   let 
-    fun split_conjunctions (Const (@{const_name "op &"}, _) $ t1 $ t2) =
+    fun split_conjunctions (Const (@{const_name HOL.conj}, _) $ t1 $ t2) =
     (split_conjunctions t1) @ (split_conjunctions t2)
     | split_conjunctions t = [t]
   in
@@ -259,7 +259,8 @@
 
 fun rewrite_intros thy =
   Simplifier.full_simplify (HOL_basic_ss addsimps [@{thm all_not_ex}])
-  #> Simplifier.full_simplify (HOL_basic_ss addsimps @{thms bool_simps} addsimps @{thms nnf_simps})
+  #> Simplifier.full_simplify
+    (HOL_basic_ss addsimps (tl @{thms bool_simps}) addsimps @{thms nnf_simps})
   #> map_term thy (maps_premises (split_conjs thy))
 
 fun print_specs options thy msg ths =
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -168,10 +168,10 @@
     mk_split_lambda' xs t
   end;
 
-fun strip_imp_prems (Const(@{const_name "op -->"}, _) $ A $ B) = A :: strip_imp_prems B
+fun strip_imp_prems (Const(@{const_name HOL.implies}, _) $ A $ B) = A :: strip_imp_prems B
   | strip_imp_prems _ = [];
 
-fun strip_imp_concl (Const(@{const_name "op -->"}, _) $ A $ B) = strip_imp_concl B
+fun strip_imp_concl (Const(@{const_name HOL.implies}, _) $ A $ B) = strip_imp_concl B
   | strip_imp_concl A = A : term;
 
 fun strip_horn A = (strip_imp_prems A, strip_imp_concl A);
@@ -185,10 +185,9 @@
 
 fun compile_term compilation options ctxt t =
   let
-    val ctxt' = ProofContext.theory (Context.copy_thy) ctxt
-    val thy = (ProofContext.theory_of ctxt') 
+    val thy = Theory.copy (ProofContext.theory_of ctxt)
     val (vs, t') = strip_abs t
-    val vs' = Variable.variant_frees ctxt' [] vs
+    val vs' = Variable.variant_frees ctxt [] vs
     val t'' = subst_bounds (map Free (rev vs'), t')
     val (prems, concl) = strip_horn t''
     val constname = "pred_compile_quickcheck"
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_specialisation.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_specialisation.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -18,8 +18,7 @@
 structure Specialisations = Theory_Data
 (
   type T = (term * term) Item_Net.T;
-  val empty = Item_Net.init ((op aconv o pairself fst) : (term * term) * (term * term) -> bool)
-    (single o fst);
+  val empty : T = Item_Net.init (op aconv o pairself fst) (single o fst);
   val extend = I;
   val merge = Item_Net.merge;
 )
--- a/src/HOL/Tools/Qelim/cooper.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Qelim/cooper.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -28,7 +28,7 @@
    @{term "op * :: int => _"}, @{term "op * :: nat => _"},
    @{term "op div :: int => _"}, @{term "op div :: nat => _"},
    @{term "op mod :: int => _"}, @{term "op mod :: nat => _"},
-   @{term "op &"}, @{term "op |"}, @{term "op -->"}, 
+   @{term HOL.conj}, @{term HOL.disj}, @{term HOL.implies}, 
    @{term "op = :: int => _"}, @{term "op = :: nat => _"}, @{term "op = :: bool => _"},
    @{term "op < :: int => _"}, @{term "op < :: nat => _"},
    @{term "op <= :: int => _"}, @{term "op <= :: nat => _"},
@@ -120,10 +120,10 @@
 
 fun whatis x ct =
 ( case (term_of ct) of
-  Const(@{const_name "op &"},_)$_$_ => And (Thm.dest_binop ct)
-| Const (@{const_name "op |"},_)$_$_ => Or (Thm.dest_binop ct)
-| Const (@{const_name "op ="},_)$y$_ => if term_of x aconv y then Eq (Thm.dest_arg ct) else Nox
-| Const (@{const_name Not},_) $ (Const (@{const_name "op ="},_)$y$_) =>
+  Const(@{const_name HOL.conj},_)$_$_ => And (Thm.dest_binop ct)
+| Const (@{const_name HOL.disj},_)$_$_ => Or (Thm.dest_binop ct)
+| Const (@{const_name HOL.eq},_)$y$_ => if term_of x aconv y then Eq (Thm.dest_arg ct) else Nox
+| Const (@{const_name Not},_) $ (Const (@{const_name HOL.eq},_)$y$_) =>
   if term_of x aconv y then NEq (funpow 2 Thm.dest_arg ct) else Nox
 | Const (@{const_name Orderings.less}, _) $ y$ z =>
    if term_of x aconv y then Lt (Thm.dest_arg ct)
@@ -274,7 +274,7 @@
   | lin vs (Const (@{const_name Not},T)$t) = Const (@{const_name Not},T)$ (lin vs t)
   | lin (vs as x::_) (Const(@{const_name Rings.dvd},_)$d$t) =
     HOLogic.mk_binrel @{const_name Rings.dvd} (numeral1 abs d, lint vs t)
-  | lin (vs as x::_) ((b as Const(@{const_name "op ="},_))$s$t) =
+  | lin (vs as x::_) ((b as Const(@{const_name HOL.eq},_))$s$t) =
      (case lint vs (subC$t$s) of
       (t as a$(m$c$y)$r) =>
         if x <> y then b$zero$t
@@ -345,7 +345,7 @@
    case (term_of t) of
     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
+      [@{const_name HOL.eq}, @{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
     then (ins (dest_number c) acc,dacc) else (acc,dacc)
   | Const(s,_)$_$(Const(@{const_name Groups.times},_)$c$y) =>
     if x aconv y andalso member (op =)
@@ -353,8 +353,8 @@
     then (ins (dest_number c) acc, dacc) else (acc,dacc)
   | 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_number c) dacc) else (acc,dacc)
-  | Const(@{const_name "op &"},_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
-  | Const(@{const_name "op |"},_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
+  | Const(@{const_name HOL.conj},_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
+  | Const(@{const_name HOL.disj},_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
   | Const (@{const_name Not},_)$_ => h (acc,dacc) (Thm.dest_arg t)
   | _ => (acc, dacc)
   val (cs,ds) = h ([],[]) p
@@ -382,12 +382,12 @@
     end
   fun unit_conv t =
    case (term_of t) of
-   Const(@{const_name "op &"},_)$_$_ => Conv.binop_conv unit_conv t
-  | Const(@{const_name "op |"},_)$_$_ => Conv.binop_conv unit_conv t
+   Const(@{const_name HOL.conj},_)$_$_ => Conv.binop_conv unit_conv t
+  | Const(@{const_name HOL.disj},_)$_$_ => Conv.binop_conv unit_conv t
   | Const (@{const_name Not},_)$_ => Conv.arg_conv unit_conv t
   | 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
+      [@{const_name HOL.eq}, @{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
     then cv (l div dest_number c) t else Thm.reflexive t
   | Const(s,_)$_$(Const(@{const_name Groups.times},_)$c$y) =>
     if x=y andalso member (op =)
@@ -569,7 +569,7 @@
 fun add_bools t =
   let
     val ops = [@{term "op = :: int => _"}, @{term "op < :: int => _"}, @{term "op <= :: int => _"},
-      @{term "op &"}, @{term "op |"}, @{term "op -->"}, @{term "op = :: bool => _"},
+      @{term HOL.conj}, @{term HOL.disj}, @{term HOL.implies}, @{term "op = :: bool => _"},
       @{term "Not"}, @{term "All :: (int => _) => _"},
       @{term "Ex :: (int => _) => _"}, @{term "True"}, @{term "False"}];
     val is_op = member (op =) ops;
@@ -612,11 +612,11 @@
 
 fun fm_of_term ps vs (Const (@{const_name True}, _)) = Proc.T
   | fm_of_term ps vs (Const (@{const_name False}, _)) = Proc.F
-  | fm_of_term ps vs (Const (@{const_name "op &"}, _) $ t1 $ t2) =
+  | fm_of_term ps vs (Const (@{const_name HOL.conj}, _) $ t1 $ t2) =
       Proc.And (fm_of_term ps vs t1, fm_of_term ps vs t2)
-  | fm_of_term ps vs (Const (@{const_name "op |"}, _) $ t1 $ t2) =
+  | fm_of_term ps vs (Const (@{const_name HOL.disj}, _) $ t1 $ t2) =
       Proc.Or (fm_of_term ps vs t1, fm_of_term ps vs t2)
-  | fm_of_term ps vs (Const (@{const_name "op -->"}, _) $ t1 $ t2) =
+  | fm_of_term ps vs (Const (@{const_name HOL.implies}, _) $ t1 $ t2) =
       Proc.Imp (fm_of_term ps vs t1, fm_of_term ps vs t2)
   | fm_of_term ps vs (@{term "op = :: bool => _ "} $ t1 $ t2) =
       Proc.Iff (fm_of_term ps vs t1, fm_of_term ps vs t2)
@@ -679,15 +679,17 @@
 
 end;
 
-val (_, oracle) = Context.>>> (Context.map_theory_result (Thm.add_oracle (Binding.name "cooper",
-  (fn (ctxt, t) => (Thm.cterm_of (ProofContext.theory_of ctxt) o Logic.mk_equals o pairself HOLogic.mk_Trueprop)
-    (t, procedure t)))));
+val (_, oracle) = Context.>>> (Context.map_theory_result
+  (Thm.add_oracle (@{binding cooper},
+    (fn (ctxt, t) =>
+      (Thm.cterm_of (ProofContext.theory_of ctxt) o Logic.mk_equals o pairself HOLogic.mk_Trueprop)
+        (t, procedure t)))));
 
 val comp_ss = HOL_ss addsimps @{thms semiring_norm};
 
 fun strip_objimp ct =
   (case Thm.term_of ct of
-    Const (@{const_name "op -->"}, _) $ _ $ _ =>
+    Const (@{const_name HOL.implies}, _) $ _ $ _ =>
       let val (A, B) = Thm.dest_binop ct
       in A :: strip_objimp B end
   | _ => [ct]);
@@ -712,7 +714,7 @@
      val qs = filter P ps
      val q = if P c then c else @{cterm "False"}
      val ng = fold_rev (fn (a,v) => fn t => Thm.capply a (Thm.cabs v t)) qvs 
-         (fold_rev (fn p => fn q => Thm.capply (Thm.capply @{cterm "op -->"} p) q) qs q)
+         (fold_rev (fn p => fn q => Thm.capply (Thm.capply @{cterm HOL.implies} p) q) qs q)
      val g = Thm.capply (Thm.capply @{cterm "op ==>"} (Thm.capply @{cterm "Trueprop"} ng)) p'
      val ntac = (case qs of [] => q aconvc @{cterm "False"}
                          | _ => false)
--- a/src/HOL/Tools/Qelim/qelim.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Qelim/qelim.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -25,8 +25,8 @@
    case (term_of p) of
     Const(s,T)$_$_ => 
        if domain_type T = HOLogic.boolT
-          andalso member (op =) [@{const_name "op &"}, @{const_name "op |"},
-            @{const_name "op -->"}, @{const_name "op ="}] s
+          andalso member (op =) [@{const_name HOL.conj}, @{const_name HOL.disj},
+            @{const_name HOL.implies}, @{const_name HOL.eq}] s
        then binop_conv (conv env) p 
        else atcv env p
   | Const(@{const_name Not},_)$_ => arg_conv (conv env) p
--- a/src/HOL/Tools/Quotient/quotient_info.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Quotient/quotient_info.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -56,10 +56,12 @@
 type maps_info = {mapfun: string, relmap: string}
 
 structure MapsData = Theory_Data
-  (type T = maps_info Symtab.table
-   val empty = Symtab.empty
-   val extend = I
-   fun merge data = Symtab.merge (K true) data)
+(
+  type T = maps_info Symtab.table
+  val empty = Symtab.empty
+  val extend = I
+  fun merge data = Symtab.merge (K true) data
+)
 
 fun maps_defined thy s =
   Symtab.defined (MapsData.get thy) s
@@ -70,7 +72,7 @@
   | NONE => raise NotFound
 
 fun maps_update_thy k minfo = MapsData.map (Symtab.update (k, minfo))
-fun maps_update k minfo = ProofContext.theory (maps_update_thy k minfo)
+fun maps_update k minfo = ProofContext.background_theory (maps_update_thy k minfo)  (* FIXME *)
 
 fun maps_attribute_aux s minfo = Thm.declaration_attribute
   (fn _ => Context.mapping (maps_update_thy s minfo) (maps_update s minfo))
@@ -120,10 +122,12 @@
 type quotdata_info = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm}
 
 structure QuotData = Theory_Data
-  (type T = quotdata_info Symtab.table
-   val empty = Symtab.empty
-   val extend = I
-   fun merge data = Symtab.merge (K true) data)
+(
+  type T = quotdata_info Symtab.table
+  val empty = Symtab.empty
+  val extend = I
+  fun merge data = Symtab.merge (K true) data
+)
 
 fun transform_quotdata phi {qtyp, rtyp, equiv_rel, equiv_thm} =
   {qtyp = Morphism.typ phi qtyp,
@@ -174,10 +178,12 @@
    for example given "nat fset" we need to find "'a fset";
    but overloaded constants share the same name *)
 structure QConstsData = Theory_Data
-  (type T = (qconsts_info list) Symtab.table
-   val empty = Symtab.empty
-   val extend = I
-   val merge = Symtab.merge_list qconsts_info_eq)
+(
+  type T = qconsts_info list Symtab.table
+  val empty = Symtab.empty
+  val extend = I
+  val merge = Symtab.merge_list qconsts_info_eq
+)
 
 fun transform_qconsts phi {qconst, rconst, def} =
   {qconst = Morphism.term phi qconst,
@@ -229,39 +235,49 @@
 
 (* equivalence relation theorems *)
 structure EquivRules = Named_Thms
-  (val name = "quot_equiv"
-   val description = "Equivalence relation theorems.")
+(
+  val name = "quot_equiv"
+  val description = "equivalence relation theorems"
+)
 
 val equiv_rules_get = EquivRules.get
 val equiv_rules_add = EquivRules.add
 
 (* respectfulness theorems *)
 structure RspRules = Named_Thms
-  (val name = "quot_respect"
-   val description = "Respectfulness theorems.")
+(
+  val name = "quot_respect"
+  val description = "respectfulness theorems"
+)
 
 val rsp_rules_get = RspRules.get
 val rsp_rules_add = RspRules.add
 
 (* preservation theorems *)
 structure PrsRules = Named_Thms
-  (val name = "quot_preserve"
-   val description = "Preservation theorems.")
+(
+  val name = "quot_preserve"
+  val description = "preservation theorems"
+)
 
 val prs_rules_get = PrsRules.get
 val prs_rules_add = PrsRules.add
 
 (* id simplification theorems *)
 structure IdSimps = Named_Thms
-  (val name = "id_simps"
-   val description = "Identity simp rules for maps.")
+(
+  val name = "id_simps"
+  val description = "identity simp rules for maps"
+)
 
 val id_simps_get = IdSimps.get
 
 (* quotient theorems *)
 structure QuotientRules = Named_Thms
-  (val name = "quot_thm"
-   val description = "Quotient theorems.")
+(
+  val name = "quot_thm"
+  val description = "quotient theorems"
+)
 
 val quotient_rules_get = QuotientRules.get
 val quotient_rules_add = QuotientRules.add
--- a/src/HOL/Tools/Quotient/quotient_tacs.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Quotient/quotient_tacs.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -12,11 +12,11 @@
   val all_injection_tac: Proof.context -> int -> tactic
   val clean_tac: Proof.context -> int -> tactic
   
-  val descend_procedure_tac: Proof.context -> int -> tactic
-  val descend_tac: Proof.context -> int -> tactic
+  val descend_procedure_tac: Proof.context -> thm list -> int -> tactic
+  val descend_tac: Proof.context -> thm list -> int -> tactic
  
-  val lift_procedure_tac: Proof.context -> thm -> int -> tactic
-  val lift_tac: Proof.context -> thm list -> int -> tactic
+  val lift_procedure_tac: Proof.context -> thm list -> thm -> int -> tactic
+  val lift_tac: Proof.context -> thm list -> thm list -> int -> tactic
 
   val lifted: Proof.context -> typ list -> thm list -> thm -> thm
   val lifted_attrib: attribute
@@ -338,7 +338,7 @@
       => rtac @{thm fun_relI} THEN' quot_true_tac ctxt unlam
 
     (* (op =) (Ball...) (Ball...) ----> (op =) (...) (...) *)
-| (Const (@{const_name "op ="},_) $
+| (Const (@{const_name HOL.eq},_) $
     (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
     (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _))
       => rtac @{thm ball_rsp} THEN' dtac @{thm QT_all}
@@ -350,7 +350,7 @@
       => rtac @{thm fun_relI} THEN' quot_true_tac ctxt unlam
 
     (* (op =) (Bex...) (Bex...) ----> (op =) (...) (...) *)
-| Const (@{const_name "op ="},_) $
+| Const (@{const_name HOL.eq},_) $
     (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
     (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
       => rtac @{thm bex_rsp} THEN' dtac @{thm QT_ex}
@@ -370,13 +370,13 @@
     (Const(@{const_name Babs},_) $ (Const (@{const_name Respects}, _) $ _) $ _))
       => rtac @{thm babs_rsp} THEN' RANGE [quotient_tac ctxt]
 
-| Const (@{const_name "op ="},_) $ (R $ _ $ _) $ (_ $ _ $ _) =>
+| Const (@{const_name HOL.eq},_) $ (R $ _ $ _) $ (_ $ _ $ _) =>
    (rtac @{thm refl} ORELSE'
     (equals_rsp_tac R ctxt THEN' RANGE [
        quot_true_tac ctxt (fst o dest_bcomb), quot_true_tac ctxt (snd o dest_bcomb)]))
 
     (* reflexivity of operators arising from Cong_tac *)
-| Const (@{const_name "op ="},_) $ _ $ _ => rtac @{thm refl}
+| Const (@{const_name HOL.eq},_) $ _ $ _ => rtac @{thm refl}
 
    (* respectfulness of constants; in particular of a simple relation *)
 | _ $ (Const _) $ (Const _)  (* fun_rel, list_rel, etc but not equality *)
@@ -606,9 +606,9 @@
   val rtrm' = HOLogic.dest_Trueprop rtrm
   val qtrm' = HOLogic.dest_Trueprop qtrm
   val reg_goal = regularize_trm_chk ctxt (rtrm', qtrm')
-    handle (LIFT_MATCH msg) => lift_match_error ctxt msg rtrm qtrm
+    handle LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm
   val inj_goal = inj_repabs_trm_chk ctxt (reg_goal, qtrm')
-    handle (LIFT_MATCH msg) => lift_match_error ctxt msg rtrm qtrm
+    handle LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm
 in
   Drule.instantiate' []
     [SOME (cterm_of thy rtrm'),
@@ -618,10 +618,21 @@
 end
 
 
+(* Since we use Ball and Bex during the lifting and descending,
+   we cannot deal with lemmas containing them, unless we unfold
+   them by default. *)
+
+val default_unfolds = @{thms Ball_def Bex_def}
+
+
 (** descending as tactic **)
 
-fun descend_procedure_tac ctxt =
-  Object_Logic.full_atomize_tac
+fun descend_procedure_tac ctxt simps =
+let
+  val ss = (mk_minimal_ss ctxt) addsimps (simps @ default_unfolds)
+in
+  full_simp_tac ss
+  THEN' Object_Logic.full_atomize_tac
   THEN' gen_frees_tac ctxt
   THEN' SUBGOAL (fn (goal, i) =>
         let
@@ -631,11 +642,12 @@
         in
           rtac rule i
         end)
+end
 
-fun descend_tac ctxt =
+fun descend_tac ctxt simps =
 let
   val mk_tac_raw =
-    descend_procedure_tac ctxt
+    descend_procedure_tac ctxt simps
     THEN' RANGE
       [Object_Logic.rulify_tac THEN' (K all_tac),
        regularize_tac ctxt,
@@ -650,15 +662,20 @@
 
 
 (* the tactic leaves three subgoals to be proved *)
-fun lift_procedure_tac ctxt rthm =
-  Object_Logic.full_atomize_tac
+fun lift_procedure_tac ctxt simps rthm =
+let
+  val ss = (mk_minimal_ss ctxt) addsimps (simps @ default_unfolds)
+in
+  full_simp_tac ss
+  THEN' Object_Logic.full_atomize_tac
   THEN' gen_frees_tac ctxt
   THEN' SUBGOAL (fn (goal, i) =>
     let
       (* full_atomize_tac contracts eta redexes,
          so we do it also in the original theorem *)
       val rthm' = 
-        rthm |> Drule.eta_contraction_rule 
+        rthm |> full_simplify ss
+             |> Drule.eta_contraction_rule 
              |> Thm.forall_intr_frees
              |> atomize_thm 
 
@@ -666,32 +683,29 @@
     in
       (rtac rule THEN' rtac rthm') i
     end)
-
+end
 
-fun lift_single_tac ctxt rthm = 
-  lift_procedure_tac ctxt rthm
+fun lift_single_tac ctxt simps rthm = 
+  lift_procedure_tac ctxt simps rthm
   THEN' RANGE
     [ regularize_tac ctxt,
       all_injection_tac ctxt,
       clean_tac ctxt ]
 
-fun lift_tac ctxt rthms =
+fun lift_tac ctxt simps rthms =
   Goal.conjunction_tac 
-  THEN' RANGE (map (lift_single_tac ctxt) rthms)
+  THEN' RANGE (map (lift_single_tac ctxt simps) rthms)
 
 
 (* automated lifting with pre-simplification of the theorems;
    for internal usage *)
 fun lifted ctxt qtys simps rthm =
 let
-  val ss = (mk_minimal_ss ctxt) addsimps simps
-  val rthm' = asm_full_simplify ss rthm
-
-  val ((_, [rthm'']), ctxt') = Variable.import true [rthm'] ctxt
-  val goal = derive_qtrm ctxt' qtys (prop_of rthm'')
+  val ((_, [rthm']), ctxt') = Variable.import true [rthm] ctxt
+  val goal = derive_qtrm ctxt' qtys (prop_of rthm')
 in
   Goal.prove ctxt' [] [] goal 
-    (K (EVERY1 [asm_full_simp_tac ss, lift_single_tac ctxt' rthm'']))
+    (K (HEADGOAL (lift_single_tac ctxt' simps rthm')))
   |> singleton (ProofContext.export ctxt' ctxt)
 end
 
--- a/src/HOL/Tools/Quotient/quotient_term.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Quotient/quotient_term.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -267,7 +267,7 @@
   map_types (Envir.subst_type ty_inst) trm
 end
 
-fun is_eq (Const (@{const_name "op ="}, _)) = true
+fun is_eq (Const (@{const_name HOL.eq}, _)) = true
   | is_eq _ = false
 
 fun mk_rel_compose (trm1, trm2) =
@@ -485,7 +485,7 @@
        end
 
   | (Const (@{const_name Ex1}, ty) $ (Abs (_, _,
-      (Const (@{const_name "op &"}, _) $ (Const (@{const_name Set.member}, _) $ _ $
+      (Const (@{const_name HOL.conj}, _) $ (Const (@{const_name Set.member}, _) $ _ $
         (Const (@{const_name Respects}, _) $ resrel)) $ (t $ _)))),
      Const (@{const_name Ex1}, ty') $ t') =>
        let
@@ -539,12 +539,12 @@
        end
 
   | (* equalities need to be replaced by appropriate equivalence relations *)
-    (Const (@{const_name "op ="}, ty), Const (@{const_name "op ="}, ty')) =>
+    (Const (@{const_name HOL.eq}, ty), Const (@{const_name HOL.eq}, ty')) =>
          if ty = ty' then rtrm
          else equiv_relation ctxt (domain_type ty, domain_type ty')
 
   | (* in this case we just check whether the given equivalence relation is correct *)
-    (rel, Const (@{const_name "op ="}, ty')) =>
+    (rel, Const (@{const_name HOL.eq}, ty')) =>
        let
          val rel_ty = fastype_of rel
          val rel' = equiv_relation_chk ctxt (domain_type rel_ty, domain_type ty')
@@ -680,7 +680,7 @@
         if T = T' then rtrm
         else mk_repabs ctxt (T, T') rtrm
 
-  | (_, Const (@{const_name "op ="}, _)) => rtrm
+  | (_, Const (@{const_name HOL.eq}, _)) => rtrm
 
   | (_, Const (_, T')) =>
       let
--- a/src/HOL/Tools/Quotient/quotient_typ.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Quotient/quotient_typ.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -77,7 +77,8 @@
   Typedef.add_typedef false NONE (qty_name, vs, mx) 
     (typedef_term rel rty lthy) NONE typedef_tac lthy
 *)
-   Local_Theory.theory_result
+(* FIXME should really use local typedef here *)
+   Local_Theory.background_theory_result
      (Typedef.add_typedef_global false NONE
        (qty_name, map (rpair dummyS) vs, mx)
          (typedef_term rel rty lthy)
--- a/src/HOL/Tools/SMT/smt_monomorph.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/smt_monomorph.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -16,7 +16,7 @@
 
 val ignored = member (op =) [
   @{const_name All}, @{const_name Ex}, @{const_name Let}, @{const_name If},
-  @{const_name "op ="}, @{const_name zero_class.zero},
+  @{const_name HOL.eq}, @{const_name zero_class.zero},
   @{const_name one_class.one}, @{const_name number_of}]
 
 fun is_const f (n, T) = not (ignored n) andalso f T
--- a/src/HOL/Tools/SMT/smt_normalize.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/smt_normalize.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -252,7 +252,7 @@
 
 fun norm_def ctxt thm =
   (case Thm.prop_of thm of
-    @{term Trueprop} $ (Const (@{const_name "op ="}, _) $ _ $ Abs _) =>
+    @{term Trueprop} $ (Const (@{const_name HOL.eq}, _) $ _ $ Abs _) =>
       norm_def ctxt (thm RS @{thm fun_cong})
   | Const (@{const_name "=="}, _) $ _ $ Abs _ =>
       norm_def ctxt (thm RS @{thm meta_eq_to_obj_eq})
--- a/src/HOL/Tools/SMT/smt_solver.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/smt_solver.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -310,18 +310,18 @@
 (* setup *)
 
 val setup =
-  Attrib.setup (Binding.name "smt_solver")
+  Attrib.setup @{binding smt_solver}
     (Scan.lift (Parse.$$$ "=" |-- Args.name) >>
       (Thm.declaration_attribute o K o select_solver))
     "SMT solver configuration" #>
   setup_timeout #>
   setup_trace #>
   setup_fixed_certificates #>
-  Attrib.setup (Binding.name "smt_certificates")
+  Attrib.setup @{binding smt_certificates}
     (Scan.lift (Parse.$$$ "=" |-- Args.name) >>
       (Thm.declaration_attribute o K o select_certificates))
     "SMT certificates" #>
-  Method.setup (Binding.name "smt") smt_method
+  Method.setup @{binding smt} smt_method
     "Applies an SMT solver to the current goal."
 
 
--- a/src/HOL/Tools/SMT/smtlib_interface.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/smtlib_interface.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -159,15 +159,15 @@
 fun conn @{const_name True} = SOME "true"
   | conn @{const_name False} = SOME "false"
   | conn @{const_name Not} = SOME "not"
-  | conn @{const_name "op &"} = SOME "and"
-  | conn @{const_name "op |"} = SOME "or"
-  | conn @{const_name "op -->"} = SOME "implies"
-  | conn @{const_name "op ="} = SOME "iff"
+  | conn @{const_name HOL.conj} = SOME "and"
+  | conn @{const_name HOL.disj} = SOME "or"
+  | conn @{const_name HOL.implies} = SOME "implies"
+  | conn @{const_name HOL.eq} = SOME "iff"
   | conn @{const_name If} = SOME "if_then_else"
   | conn _ = NONE
 
 fun pred @{const_name distinct} _ = SOME "distinct"
-  | pred @{const_name "op ="} _ = SOME "="
+  | pred @{const_name HOL.eq} _ = SOME "="
   | pred @{const_name term_eq} _ = SOME "="
   | pred @{const_name less} T = if_int_type T "<"
   | pred @{const_name less_eq} T = if_int_type T "<="
--- a/src/HOL/Tools/SMT/z3_interface.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_interface.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -170,13 +170,13 @@
 val mk_true = @{cterm "~False"}
 val mk_false = @{cterm False}
 val mk_not = Thm.capply @{cterm Not}
-val mk_implies = Thm.mk_binop @{cterm "op -->"}
+val mk_implies = Thm.mk_binop @{cterm HOL.implies}
 val mk_iff = Thm.mk_binop @{cterm "op = :: bool => _"}
 
 fun mk_nary _ cu [] = cu
   | mk_nary ct _ cts = uncurry (fold_rev (Thm.mk_binop ct)) (split_last cts)
 
-val eq = mk_inst_pair destT1 @{cpat "op ="}
+val eq = mk_inst_pair destT1 @{cpat HOL.eq}
 fun mk_eq ct cu = Thm.mk_binop (instT' ct eq) ct cu
 
 val if_term = mk_inst_pair (destT1 o destT2) @{cpat If}
@@ -216,8 +216,8 @@
     (Sym ("true", _), []) => SOME mk_true
   | (Sym ("false", _), []) => SOME mk_false
   | (Sym ("not", _), [ct]) => SOME (mk_not ct)
-  | (Sym ("and", _), _) => SOME (mk_nary @{cterm "op &"} mk_true cts)
-  | (Sym ("or", _), _) => SOME (mk_nary @{cterm "op |"} mk_false cts)
+  | (Sym ("and", _), _) => SOME (mk_nary @{cterm HOL.conj} mk_true cts)
+  | (Sym ("or", _), _) => SOME (mk_nary @{cterm HOL.disj} mk_false cts)
   | (Sym ("implies", _), [ct, cu]) => SOME (mk_implies ct cu)
   | (Sym ("iff", _), [ct, cu]) => SOME (mk_iff ct cu)
   | (Sym ("~", _), [ct, cu]) => SOME (mk_iff ct cu)
--- a/src/HOL/Tools/SMT/z3_proof_literals.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_proof_literals.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -62,14 +62,14 @@
 val is_neg = (fn @{term Not} $ _ => true | _ => false)
 fun is_neg' f = (fn @{term Not} $ t => f t | _ => false)
 val is_dneg = is_neg' is_neg
-val is_conj = (fn @{term "op &"} $ _ $ _ => true | _ => false)
-val is_disj = (fn @{term "op |"} $ _ $ _ => true | _ => false)
+val is_conj = (fn @{term HOL.conj} $ _ $ _ => true | _ => false)
+val is_disj = (fn @{term HOL.disj} $ _ $ _ => true | _ => false)
 
 fun dest_disj_term' f = (fn
-    @{term Not} $ (@{term "op |"} $ t $ u) => SOME (f t, f u)
+    @{term Not} $ (@{term HOL.disj} $ t $ u) => SOME (f t, f u)
   | _ => NONE)
 
-val dest_conj_term = (fn @{term "op &"} $ t $ u => SOME (t, u) | _ => NONE)
+val dest_conj_term = (fn @{term HOL.conj} $ t $ u => SOME (t, u) | _ => NONE)
 val dest_disj_term =
   dest_disj_term' (fn @{term Not} $ t => t | t => @{term Not} $ t)
 
@@ -202,11 +202,11 @@
     | comp_disj ((true, thm1), (false, thm2)) = comp2 disj3 thm1 thm2
     | comp_disj ((true, thm1), (true, thm2)) = comp2 disj4 thm1 thm2
 
-  fun dest_conj (@{term "op &"} $ t $ u) = ((false, t), (false, u))
+  fun dest_conj (@{term HOL.conj} $ t $ u) = ((false, t), (false, u))
     | dest_conj t = raise TERM ("dest_conj", [t])
 
   val neg = (fn @{term Not} $ t => (true, t) | t => (false, @{term Not} $ t))
-  fun dest_disj (@{term Not} $ (@{term "op |"} $ t $ u)) = (neg t, neg u)
+  fun dest_disj (@{term Not} $ (@{term HOL.disj} $ t $ u)) = (neg t, neg u)
     | dest_disj t = raise TERM ("dest_disj", [t])
 
   val dnegE = T.precompose (single o d2 o d1) @{thm notnotD}
@@ -234,7 +234,7 @@
         @{term Not} $ (@{term Not} $ t) => (T.compose dnegI, lookup t)
       | @{term Not} $ (@{term "op = :: bool => _"} $ t $ (@{term Not} $ u)) =>
           (T.compose negIffI, lookup (iff_const $ u $ t))
-      | @{term Not} $ ((eq as Const (@{const_name "op ="}, _)) $ t $ u) =>
+      | @{term Not} $ ((eq as Const (@{const_name HOL.eq}, _)) $ t $ u) =>
           let fun rewr lit = lit COMP @{thm not_sym}
           in (rewr, lookup (@{term Not} $ (eq $ u $ t))) end
       | _ =>
--- a/src/HOL/Tools/SMT/z3_proof_reconstruction.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_proof_reconstruction.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -298,13 +298,13 @@
     let val (ct1, ct2) = Thm.dest_binop (Thm.dest_arg ct)
     in
       (case Thm.term_of ct1 of
-        @{term Not} $ (@{term "op &"} $ _ $ _) =>
+        @{term Not} $ (@{term HOL.conj} $ _ $ _) =>
           prove disjI1 (Thm.dest_arg ct1, true) (ct2, true)
-      | @{term "op &"} $ _ $ _ =>
+      | @{term HOL.conj} $ _ $ _ =>
           prove disjI3 (Thm.capply @{cterm Not} ct2, false) (ct1, true)
-      | @{term Not} $ (@{term "op |"} $ _ $ _) =>
+      | @{term Not} $ (@{term HOL.disj} $ _ $ _) =>
           prove disjI3 (Thm.capply @{cterm Not} ct2, false) (ct1, false)
-      | @{term "op |"} $ _ $ _ =>
+      | @{term HOL.disj} $ _ $ _ =>
           prove disjI2 (Thm.capply @{cterm Not} ct1, false) (ct2, true)
       | Const (@{const_name distinct}, _) $ _ =>
           let
@@ -477,8 +477,8 @@
 
   fun mono f (cp as (cl, _)) =
     (case Term.head_of (Thm.term_of cl) of
-      @{term "op &"} => prove_nary L.is_conj (prove_eq_exn f)
-    | @{term "op |"} => prove_nary L.is_disj (prove_eq_exn f)
+      @{term HOL.conj} => prove_nary L.is_conj (prove_eq_exn f)
+    | @{term HOL.disj} => prove_nary L.is_disj (prove_eq_exn f)
     | Const (@{const_name distinct}, _) => prove_distinct (prove_eq_safe f)
     | _ => prove (prove_eq_safe f)) cp
 in
--- a/src/HOL/Tools/SMT/z3_proof_tools.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_proof_tools.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -196,10 +196,10 @@
       | @{term True} => pair ct
       | @{term False} => pair ct
       | @{term Not} $ _ => abstr1 cvs ct
-      | @{term "op &"} $ _ $ _ => abstr2 cvs ct
-      | @{term "op |"} $ _ $ _ => abstr2 cvs ct
-      | @{term "op -->"} $ _ $ _ => abstr2 cvs ct
-      | Const (@{const_name "op ="}, _) $ _ $ _ => abstr2 cvs ct
+      | @{term HOL.conj} $ _ $ _ => abstr2 cvs ct
+      | @{term HOL.disj} $ _ $ _ => abstr2 cvs ct
+      | @{term HOL.implies} $ _ $ _ => abstr2 cvs ct
+      | Const (@{const_name HOL.eq}, _) $ _ $ _ => abstr2 cvs ct
       | Const (@{const_name distinct}, _) $ _ =>
           if ext_logic then abs_arg (abs_list abstr fresh_abstraction) cvs ct
           else fresh_abstraction cvs ct
--- a/src/HOL/Tools/Sledgehammer/clausifier.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/clausifier.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -68,8 +68,8 @@
         (*Universal quant: insert a free variable into body and continue*)
         let val fname = Name.variant (OldTerm.add_term_names (p,[])) a
         in dec_sko (subst_bound (Free(fname,T), p)) rhss end
-      | dec_sko (@{const "op &"} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q
-      | dec_sko (@{const "op |"} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q
+      | dec_sko (@{const HOL.conj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q
+      | dec_sko (@{const HOL.disj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q
       | dec_sko (@{const Trueprop} $ p) rhss = dec_sko p rhss
       | dec_sko _ rhss = rhss
   in  dec_sko (prop_of th) []  end;
@@ -83,7 +83,7 @@
    (Cf. "extensionalize_term" in "Sledgehammer_Translate".) *)
 fun extensionalize_theorem th =
   case prop_of th of
-    _ $ (Const (@{const_name "op ="}, Type (_, [Type (@{type_name fun}, _), _]))
+    _ $ (Const (@{const_name HOL.eq}, Type (_, [Type (@{type_name fun}, _), _]))
          $ _ $ Abs (s, _, _)) => extensionalize_theorem (th RS fun_cong_all)
   | _ => th
 
--- a/src/HOL/Tools/Sledgehammer/metis_clauses.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/metis_clauses.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -38,11 +38,10 @@
   val const_prefix: string
   val type_const_prefix: string
   val class_prefix: string
-  val union_all: ''a list list -> ''a list
   val invert_const: string -> string
   val ascii_of: string -> string
-  val undo_ascii_of: string -> string
-  val strip_prefix_and_undo_ascii: string -> string -> string option
+  val unascii_of: string -> string
+  val strip_prefix_and_unascii: string -> string -> string option
   val make_bound_var : string -> string
   val make_schematic_var : string * int -> string
   val make_fixed_var : string -> string
@@ -95,10 +94,10 @@
 val const_trans_table =
   Symtab.make [(@{type_name Product_Type.prod}, "prod"),
                (@{type_name Sum_Type.sum}, "sum"),
-               (@{const_name "op ="}, "equal"),
-               (@{const_name "op &"}, "and"),
-               (@{const_name "op |"}, "or"),
-               (@{const_name "op -->"}, "implies"),
+               (@{const_name HOL.eq}, "equal"),
+               (@{const_name HOL.conj}, "and"),
+               (@{const_name HOL.disj}, "or"),
+               (@{const_name HOL.implies}, "implies"),
                (@{const_name Set.member}, "member"),
                (@{const_name fequal}, "fequal"),
                (@{const_name COMBI}, "COMBI"),
@@ -112,7 +111,7 @@
 
 (* Invert the table of translations between Isabelle and ATPs. *)
 val const_trans_table_inv =
-  Symtab.update ("fequal", @{const_name "op ="})
+  Symtab.update ("fequal", @{const_name HOL.eq})
                 (Symtab.make (map swap (Symtab.dest const_trans_table)))
 
 val invert_const = perhaps (Symtab.lookup const_trans_table_inv)
@@ -121,7 +120,7 @@
   Alphanumeric characters are left unchanged.
   The character _ goes to __
   Characters in the range ASCII space to / go to _A to _P, respectively.
-  Other printing characters go to _nnn where nnn is the decimal ASCII code.*)
+  Other characters go to _nnn where nnn is the decimal ASCII code.*)
 val A_minus_space = Char.ord #"A" - Char.ord #" ";
 
 fun stringN_of_int 0 _ = ""
@@ -132,9 +131,7 @@
   else if c = #"_" then "__"
   else if #" " <= c andalso c <= #"/"
        then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
-  else if Char.isPrint c
-       then ("_" ^ stringN_of_int 3 (Char.ord c))  (*fixed width, in case more digits follow*)
-  else ""
+  else ("_" ^ stringN_of_int 3 (Char.ord c))  (*fixed width, in case more digits follow*)
 
 val ascii_of = String.translate ascii_of_c;
 
@@ -142,29 +139,28 @@
 
 (*We don't raise error exceptions because this code can run inside the watcher.
   Also, the errors are "impossible" (hah!)*)
-fun undo_ascii_aux rcs [] = String.implode(rev rcs)
-  | undo_ascii_aux rcs [#"_"] = undo_ascii_aux (#"_"::rcs) []  (*ERROR*)
+fun unascii_aux rcs [] = String.implode(rev rcs)
+  | unascii_aux rcs [#"_"] = unascii_aux (#"_"::rcs) []  (*ERROR*)
       (*Three types of _ escapes: __, _A to _P, _nnn*)
-  | undo_ascii_aux rcs (#"_" :: #"_" :: cs) = undo_ascii_aux (#"_"::rcs) cs
-  | undo_ascii_aux rcs (#"_" :: c :: cs) =
+  | unascii_aux rcs (#"_" :: #"_" :: cs) = unascii_aux (#"_"::rcs) cs
+  | unascii_aux rcs (#"_" :: c :: cs) =
       if #"A" <= c andalso c<= #"P"  (*translation of #" " to #"/"*)
-      then undo_ascii_aux (Char.chr(Char.ord c - A_minus_space) :: rcs) cs
+      then unascii_aux (Char.chr(Char.ord c - A_minus_space) :: rcs) cs
       else
         let val digits = List.take (c::cs, 3) handle Subscript => []
         in
             case Int.fromString (String.implode digits) of
-                NONE => undo_ascii_aux (c:: #"_"::rcs) cs  (*ERROR*)
-              | SOME n => undo_ascii_aux (Char.chr n :: rcs) (List.drop (cs, 2))
+                NONE => unascii_aux (c:: #"_"::rcs) cs  (*ERROR*)
+              | SOME n => unascii_aux (Char.chr n :: rcs) (List.drop (cs, 2))
         end
-  | undo_ascii_aux rcs (c::cs) = undo_ascii_aux (c::rcs) cs;
-
-val undo_ascii_of = undo_ascii_aux [] o String.explode;
+  | unascii_aux rcs (c::cs) = unascii_aux (c::rcs) cs
+val unascii_of = unascii_aux [] o String.explode
 
 (* If string s has the prefix s1, return the result of deleting it,
    un-ASCII'd. *)
-fun strip_prefix_and_undo_ascii s1 s =
+fun strip_prefix_and_unascii s1 s =
   if String.isPrefix s1 s then
-    SOME (undo_ascii_of (String.extract (s, size s1, NONE)))
+    SOME (unascii_of (String.extract (s, size s1, NONE)))
   else
     NONE
 
@@ -189,8 +185,8 @@
     SOME c' => c'
   | NONE => ascii_of c
 
-(* "op =" MUST BE "equal" because it's built into ATPs. *)
-fun make_fixed_const @{const_name "op ="} = "equal"
+(* HOL.eq MUST BE "equal" because it's built into ATPs. *)
+fun make_fixed_const @{const_name HOL.eq} = "equal"
   | make_fixed_const c = const_prefix ^ lookup_const c
 
 fun make_fixed_type_const c = type_const_prefix ^ lookup_const c
@@ -434,7 +430,7 @@
 
 fun literals_of_term1 args thy (@{const Trueprop} $ P) =
     literals_of_term1 args thy P
-  | literals_of_term1 args thy (@{const "op |"} $ P $ Q) =
+  | literals_of_term1 args thy (@{const HOL.disj} $ P $ Q) =
     literals_of_term1 (literals_of_term1 args thy P) thy Q
   | literals_of_term1 (lits, ts) thy P =
     let val ((pred, ts'), pol) = predicate_of thy (P, true) in
@@ -514,8 +510,8 @@
 (*Type constructors used to instantiate overloaded constants are the only ones needed.*)
 fun add_type_consts_in_term thy =
   let
-    val const_typargs = Sign.const_typargs thy
-    fun aux (Const x) = fold (fold_type_consts set_insert) (const_typargs x)
+    fun aux (Const x) =
+        fold (fold_type_consts set_insert) (Sign.const_typargs thy x)
       | aux (Abs (_, _, u)) = aux u
       | aux (Const (@{const_name skolem_id}, _) $ _) = I
       | aux (t $ u) = aux t #> aux u
--- a/src/HOL/Tools/Sledgehammer/metis_tactics.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/metis_tactics.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -224,19 +224,19 @@
 
 fun make_tvar s = TVar (("'" ^ s, 0), HOLogic.typeS)
 
-fun smart_invert_const "fequal" = @{const_name "op ="}
+fun smart_invert_const "fequal" = @{const_name HOL.eq}
   | smart_invert_const s = invert_const s
 
 fun hol_type_from_metis_term _ (Metis.Term.Var v) =
-     (case strip_prefix_and_undo_ascii tvar_prefix v of
+     (case strip_prefix_and_unascii tvar_prefix v of
           SOME w => make_tvar w
         | NONE   => make_tvar v)
   | hol_type_from_metis_term ctxt (Metis.Term.Fn(x, tys)) =
-     (case strip_prefix_and_undo_ascii type_const_prefix x of
+     (case strip_prefix_and_unascii type_const_prefix x of
           SOME tc => Term.Type (smart_invert_const tc,
                                 map (hol_type_from_metis_term ctxt) tys)
         | NONE    =>
-      case strip_prefix_and_undo_ascii tfree_prefix x of
+      case strip_prefix_and_unascii tfree_prefix x of
           SOME tf => mk_tfree ctxt tf
         | NONE    => raise Fail ("hol_type_from_metis_term: " ^ x));
 
@@ -246,10 +246,10 @@
       val _ = trace_msg (fn () => "hol_term_from_metis_PT: " ^
                                   Metis.Term.toString fol_tm)
       fun tm_to_tt (Metis.Term.Var v) =
-             (case strip_prefix_and_undo_ascii tvar_prefix v of
+             (case strip_prefix_and_unascii tvar_prefix v of
                   SOME w => Type (make_tvar w)
                 | NONE =>
-              case strip_prefix_and_undo_ascii schematic_var_prefix v of
+              case strip_prefix_and_unascii schematic_var_prefix v of
                   SOME w => Term (mk_var (w, HOLogic.typeT))
                 | NONE   => Term (mk_var (v, HOLogic.typeT)) )
                     (*Var from Metis with a name like _nnn; possibly a type variable*)
@@ -264,9 +264,9 @@
             end
         | tm_to_tt (Metis.Term.Fn (fname, args)) = applic_to_tt (fname,args)
       and applic_to_tt ("=",ts) =
-            Term (list_comb(Const (@{const_name "op ="}, HOLogic.typeT), terms_of (map tm_to_tt ts)))
+            Term (list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), terms_of (map tm_to_tt ts)))
         | applic_to_tt (a,ts) =
-            case strip_prefix_and_undo_ascii const_prefix a of
+            case strip_prefix_and_unascii const_prefix a of
                 SOME b =>
                   let val c = smart_invert_const b
                       val ntypes = num_type_args thy c
@@ -284,14 +284,14 @@
                                    cat_lines (map (Syntax.string_of_term ctxt) (terms_of tts)))
                      end
               | NONE => (*Not a constant. Is it a type constructor?*)
-            case strip_prefix_and_undo_ascii type_const_prefix a of
+            case strip_prefix_and_unascii type_const_prefix a of
                 SOME b =>
                   Type (Term.Type (smart_invert_const b, types_of (map tm_to_tt ts)))
               | NONE => (*Maybe a TFree. Should then check that ts=[].*)
-            case strip_prefix_and_undo_ascii tfree_prefix a of
+            case strip_prefix_and_unascii tfree_prefix a of
                 SOME b => Type (mk_tfree ctxt b)
               | NONE => (*a fixed variable? They are Skolem functions.*)
-            case strip_prefix_and_undo_ascii fixed_var_prefix a of
+            case strip_prefix_and_unascii fixed_var_prefix a of
                 SOME b =>
                   let val opr = Term.Free(b, HOLogic.typeT)
                   in  apply_list opr (length ts) (map tm_to_tt ts)  end
@@ -307,16 +307,16 @@
   let val _ = trace_msg (fn () => "hol_term_from_metis_FT: " ^
                                   Metis.Term.toString fol_tm)
       fun cvt (Metis.Term.Fn ("ti", [Metis.Term.Var v, _])) =
-             (case strip_prefix_and_undo_ascii schematic_var_prefix v of
+             (case strip_prefix_and_unascii schematic_var_prefix v of
                   SOME w =>  mk_var(w, dummyT)
                 | NONE   => mk_var(v, dummyT))
         | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn ("=",[]), _])) =
-            Const (@{const_name "op ="}, HOLogic.typeT)
+            Const (@{const_name HOL.eq}, HOLogic.typeT)
         | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (x,[]), ty])) =
-           (case strip_prefix_and_undo_ascii const_prefix x of
+           (case strip_prefix_and_unascii const_prefix x of
                 SOME c => Const (smart_invert_const c, dummyT)
               | NONE => (*Not a constant. Is it a fixed variable??*)
-            case strip_prefix_and_undo_ascii fixed_var_prefix x of
+            case strip_prefix_and_unascii fixed_var_prefix x of
                 SOME v => Free (v, hol_type_from_metis_term ctxt ty)
               | NONE => raise Fail ("hol_term_from_metis_FT bad constant: " ^ x))
         | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (".",[tm1,tm2]), _])) =
@@ -325,12 +325,12 @@
             cvt tm1 $ cvt tm2
         | cvt (Metis.Term.Fn ("{}", [arg])) = cvt arg   (*hBOOL*)
         | cvt (Metis.Term.Fn ("=", [tm1,tm2])) =
-            list_comb(Const (@{const_name "op ="}, HOLogic.typeT), map cvt [tm1,tm2])
+            list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), map cvt [tm1,tm2])
         | cvt (t as Metis.Term.Fn (x, [])) =
-           (case strip_prefix_and_undo_ascii const_prefix x of
+           (case strip_prefix_and_unascii const_prefix x of
                 SOME c => Const (smart_invert_const c, dummyT)
               | NONE => (*Not a constant. Is it a fixed variable??*)
-            case strip_prefix_and_undo_ascii fixed_var_prefix x of
+            case strip_prefix_and_unascii fixed_var_prefix x of
                 SOME v => Free (v, dummyT)
               | NONE => (trace_msg (fn () => "hol_term_from_metis_FT bad const: " ^ x);
                   hol_term_from_metis_PT ctxt t))
@@ -410,11 +410,11 @@
                                        " in " ^ Display.string_of_thm ctxt i_th);
                  NONE)
       fun remove_typeinst (a, t) =
-            case strip_prefix_and_undo_ascii schematic_var_prefix a of
+            case strip_prefix_and_unascii schematic_var_prefix a of
                 SOME b => SOME (b, t)
-              | NONE   => case strip_prefix_and_undo_ascii tvar_prefix a of
+              | NONE => case strip_prefix_and_unascii tvar_prefix a of
                 SOME _ => NONE          (*type instantiations are forbidden!*)
-              | NONE   => SOME (a,t)    (*internal Metis var?*)
+              | NONE => SOME (a,t)    (*internal Metis var?*)
       val _ = trace_msg (fn () => "  isa th: " ^ Display.string_of_thm ctxt i_th)
       val substs = map_filter remove_typeinst (Metis.Subst.toList fsubst)
       val (vars,rawtms) = ListPair.unzip (map_filter subst_translation substs)
@@ -480,7 +480,7 @@
       val c_t = cterm_incr_types thy refl_idx i_t
   in  cterm_instantiate [(refl_x, c_t)] REFL_THM  end;
 
-fun get_ty_arg_size _ (Const (@{const_name "op ="}, _)) = 0  (*equality has no type arguments*)
+fun get_ty_arg_size _ (Const (@{const_name HOL.eq}, _)) = 0  (*equality has no type arguments*)
   | get_ty_arg_size thy (Const (c, _)) = (num_type_args thy c handle TYPE _ => 0)
   | get_ty_arg_size _ _ = 0;
 
@@ -529,13 +529,13 @@
                       " isa-term: " ^  Syntax.string_of_term ctxt tm ^
                       " fol-term: " ^ Metis.Term.toString t)
       fun path_finder FO tm ps _ = path_finder_FO tm ps
-        | path_finder HO (tm as Const(@{const_name "op ="},_) $ _ $ _) (p::ps) _ =
+        | path_finder HO (tm as Const(@{const_name HOL.eq},_) $ _ $ _) (p::ps) _ =
              (*equality: not curried, as other predicates are*)
              if p=0 then path_finder_HO tm (0::1::ps)  (*select first operand*)
              else path_finder_HO tm (p::ps)        (*1 selects second operand*)
         | path_finder HO tm (_ :: ps) (Metis.Term.Fn ("{}", [_])) =
              path_finder_HO tm ps      (*if not equality, ignore head to skip hBOOL*)
-        | path_finder FT (tm as Const(@{const_name "op ="}, _) $ _ $ _) (p::ps)
+        | path_finder FT (tm as Const(@{const_name HOL.eq}, _) $ _ $ _) (p::ps)
                             (Metis.Term.Fn ("=", [t1,t2])) =
              (*equality: not curried, as other predicates are*)
              if p=0 then path_finder_FT tm (0::1::ps)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -9,6 +9,7 @@
 signature SLEDGEHAMMER =
 sig
   type failure = ATP_Systems.failure
+  type locality = Sledgehammer_Fact_Filter.locality
   type relevance_override = Sledgehammer_Fact_Filter.relevance_override
   type minimize_command = Sledgehammer_Proof_Reconstruct.minimize_command
   type params =
@@ -18,11 +19,9 @@
      atps: string list,
      full_types: bool,
      explicit_apply: bool,
-     relevance_threshold: real,
-     relevance_convergence: real,
-     max_relevant_per_iter: int option,
+     relevance_thresholds: real * real,
+     max_relevant: int option,
      theory_relevant: bool option,
-     defs_relevant: bool,
      isar_proof: bool,
      isar_shrink_factor: int,
      timeout: Time.time}
@@ -30,16 +29,16 @@
     {subgoal: int,
      goal: Proof.context * (thm list * thm),
      relevance_override: relevance_override,
-     axioms: ((string * bool) * thm) list option}
+     axioms: ((string * locality) * thm) list option}
   type prover_result =
     {outcome: failure option,
      message: string,
      pool: string Symtab.table,
-     used_thm_names: (string * bool) list,
+     used_thm_names: (string * locality) list,
      atp_run_time_in_msecs: int,
      output: string,
      proof: string,
-     axiom_names: (string * bool) vector,
+     axiom_names: (string * locality) list vector,
      conjecture_shape: int list list}
   type prover = params -> minimize_command -> problem -> prover_result
 
@@ -87,11 +86,9 @@
    atps: string list,
    full_types: bool,
    explicit_apply: bool,
-   relevance_threshold: real,
-   relevance_convergence: real,
-   max_relevant_per_iter: int option,
+   relevance_thresholds: real * real,
+   max_relevant: int option,
    theory_relevant: bool option,
-   defs_relevant: bool,
    isar_proof: bool,
    isar_shrink_factor: int,
    timeout: Time.time}
@@ -100,17 +97,17 @@
   {subgoal: int,
    goal: Proof.context * (thm list * thm),
    relevance_override: relevance_override,
-   axioms: ((string * bool) * thm) list option}
+   axioms: ((string * locality) * thm) list option}
 
 type prover_result =
   {outcome: failure option,
    message: string,
    pool: string Symtab.table,
-   used_thm_names: (string * bool) list,
+   used_thm_names: (string * locality) list,
    atp_run_time_in_msecs: int,
    output: string,
    proof: string,
-   axiom_names: (string * bool) vector,
+   axiom_names: (string * locality) list vector,
    conjecture_shape: int list list}
 
 type prover = params -> minimize_command -> problem -> prover_result
@@ -174,11 +171,11 @@
   Scan.this_string set_ClauseFormulaRelationN |-- $$ "("
   |-- Scan.repeat parse_clause_formula_pair
 val extract_clause_formula_relation =
-  Substring.full
-  #> Substring.position set_ClauseFormulaRelationN
-  #> snd #> Substring.string #> strip_spaces #> explode
-  #> parse_clause_formula_relation #> fst
+  Substring.full #> Substring.position set_ClauseFormulaRelationN
+  #> snd #> Substring.string #> strip_spaces_except_between_ident_chars
+  #> explode #> parse_clause_formula_relation #> fst
 
+(* TODO: move to "Sledgehammer_Proof_Reconstruct.ML" *)
 fun repair_conjecture_shape_and_theorem_names output conjecture_shape
                                               axiom_names =
   if String.isSubstring set_ClauseFormulaRelationN output then
@@ -190,20 +187,20 @@
       val seq = extract_clause_sequence output
       val name_map = extract_clause_formula_relation output
       fun renumber_conjecture j =
-        conjecture_prefix ^ Int.toString (j - j0)
+        conjecture_prefix ^ string_of_int (j - j0)
         |> AList.find (fn (s, ss) => member (op =) ss s) name_map
         |> map (fn s => find_index (curry (op =) s) seq + 1)
-      fun name_for_number j =
-        let
-          val axioms =
-            j |> AList.lookup (op =) name_map
-              |> these |> map_filter (try (unprefix axiom_prefix))
-              |> map undo_ascii_of
-          val chained = forall (is_true_for axiom_names) axioms
-        in (axioms |> space_implode " ", chained) end
+      fun names_for_number j =
+        j |> AList.lookup (op =) name_map |> these
+          |> map_filter (try (unprefix axiom_prefix)) |> map unascii_of
+          |> map (fn name =>
+                     (name, name |> find_first_in_list_vector axiom_names
+                                 |> the)
+                     handle Option.Option =>
+                            error ("No such fact: " ^ quote name ^ "."))
     in
       (conjecture_shape |> map (maps renumber_conjecture),
-       seq |> map name_for_number |> Vector.fromList)
+       seq |> map names_for_number |> Vector.fromList)
     end
   else
     (conjecture_shape, axiom_names)
@@ -213,20 +210,18 @@
 
 fun prover_fun atp_name
         {exec, required_execs, arguments, has_incomplete_mode, proof_delims,
-         known_failures, default_max_relevant_per_iter, default_theory_relevant,
+         known_failures, default_max_relevant, default_theory_relevant,
          explicit_forall, use_conjecture_for_hypotheses}
         ({debug, verbose, overlord, full_types, explicit_apply,
-          relevance_threshold, relevance_convergence,
-          max_relevant_per_iter, theory_relevant,
-          defs_relevant, isar_proof, isar_shrink_factor, timeout, ...} : params)
+          relevance_thresholds, max_relevant, theory_relevant, isar_proof,
+          isar_shrink_factor, timeout, ...} : params)
         minimize_command
-        ({subgoal, goal, relevance_override, axioms} : problem) =
+        ({subgoal, goal = (ctxt, (chained_ths, th)), relevance_override,
+          axioms} : problem) =
   let
-    val (ctxt, (_, th)) = goal;
-    val thy = ProofContext.theory_of ctxt
-    val (params, hyp_ts, concl_t) = strip_subgoal th subgoal
+    val (_, hyp_ts, concl_t) = strip_subgoal th subgoal
 
-    fun print s = (priority s; if debug then tracing s else ())
+    val print = priority
     fun print_v f = () |> verbose ? print o f
     fun print_d f = () |> debug ? print o f
 
@@ -234,15 +229,13 @@
       case axioms of
         SOME axioms => axioms
       | NONE =>
-        (relevant_facts full_types relevance_threshold relevance_convergence
-                        defs_relevant
-                        (the_default default_max_relevant_per_iter
-                                     max_relevant_per_iter)
+        (relevant_facts ctxt full_types relevance_thresholds
+                        (the_default default_max_relevant max_relevant)
                         (the_default default_theory_relevant theory_relevant)
-                        relevance_override goal hyp_ts concl_t
+                        relevance_override chained_ths hyp_ts concl_t
          |> tap ((fn n => print_v (fn () =>
-                      "Selected " ^ string_of_int n ^ " fact" ^ plural_s n ^
-                      " for " ^ quote atp_name ^ ".")) o length))
+                          "Selected " ^ string_of_int n ^ " fact" ^ plural_s n ^
+                          " for " ^ quote atp_name ^ ".")) o length))
 
     (* path to unique problem file *)
     val the_dest_dir = if overlord then getenv "ISABELLE_HOME_USER"
@@ -258,9 +251,10 @@
         if the_dest_dir = "" then File.tmp_path probfile
         else if File.exists (Path.explode the_dest_dir)
         then Path.append (Path.explode the_dest_dir) probfile
-        else error ("No such directory: " ^ the_dest_dir ^ ".")
+        else error ("No such directory: " ^ quote the_dest_dir ^ ".")
       end;
 
+    val measure_run_time = verbose orelse Config.get ctxt measure_runtime
     val command = Path.explode (getenv (fst exec) ^ "/" ^ snd exec)
     (* write out problem file and call prover *)
     fun command_line complete timeout probfile =
@@ -268,10 +262,8 @@
         val core = File.shell_path command ^ " " ^ arguments complete timeout ^
                    " " ^ File.shell_path probfile
       in
-        (if Config.get ctxt measure_runtime then
-           "TIMEFORMAT='%3U'; { time " ^ core ^ " ; }"
-         else
-           "exec " ^ core) ^ " 2>&1"
+        (if measure_run_time then "TIMEFORMAT='%3U'; { time " ^ core ^ " ; }"
+         else "exec " ^ core) ^ " 2>&1"
       end
     fun split_time s =
       let
@@ -300,14 +292,11 @@
                          prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n")
                        else
                          I)
-                  |>> (if Config.get ctxt measure_runtime then split_time
-                       else rpair 0)
+                  |>> (if measure_run_time then split_time else rpair 0)
                 val (proof, outcome) =
                   extract_proof_and_outcome complete res_code proof_delims
                                             known_failures output
               in (output, msecs, proof, outcome) end
-            val _ = print_d (fn () => "Preparing problem for " ^
-                                      quote atp_name ^ "...")
             val readable_names = debug andalso overlord
             val (problem, pool, conjecture_offset, axiom_names) =
               prepare_problem ctxt readable_names explicit_forall full_types
@@ -358,6 +347,11 @@
         proof_text isar_proof
             (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
             (full_types, minimize_command, proof, axiom_names, th, subgoal)
+        |>> (fn message =>
+                message ^ (if verbose then
+                             "\nATP CPU time: " ^ string_of_int msecs ^ " ms."
+                           else
+                             ""))
       | SOME failure => (string_for_failure failure, [])
   in
     {outcome = outcome, message = message, pool = pool,
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -5,19 +5,34 @@
 
 signature SLEDGEHAMMER_FACT_FILTER =
 sig
+  datatype locality = General | Intro | Elim | Simp | Local | Chained
+
   type relevance_override =
     {add: Facts.ref list,
      del: Facts.ref list,
      only: bool}
 
   val trace : bool Unsynchronized.ref
-  val name_thms_pair_from_ref :
+  val worse_irrel_freq : real Unsynchronized.ref
+  val higher_order_irrel_weight : real Unsynchronized.ref
+  val abs_rel_weight : real Unsynchronized.ref
+  val abs_irrel_weight : real Unsynchronized.ref
+  val skolem_irrel_weight : real Unsynchronized.ref
+  val intro_bonus : real Unsynchronized.ref
+  val elim_bonus : real Unsynchronized.ref
+  val simp_bonus : real Unsynchronized.ref
+  val local_bonus : real Unsynchronized.ref
+  val chained_bonus : real Unsynchronized.ref
+  val max_imperfect : real Unsynchronized.ref
+  val max_imperfect_exp : real Unsynchronized.ref
+  val threshold_divisor : real Unsynchronized.ref
+  val ridiculous_threshold : real Unsynchronized.ref
+  val name_thm_pairs_from_ref :
     Proof.context -> unit Symtab.table -> thm list -> Facts.ref
-    -> (unit -> string * bool) * thm list
+    -> ((string * locality) * thm) list
   val relevant_facts :
-    bool -> real -> real -> bool -> int -> bool -> relevance_override
-    -> Proof.context * (thm list * 'a) -> term list -> term
-    -> ((string * bool) * thm) list
+    Proof.context -> bool -> real * real -> int -> bool -> relevance_override
+    -> thm list -> term list -> term -> ((string * locality) * thm) list
 end;
 
 structure Sledgehammer_Fact_Filter : SLEDGEHAMMER_FACT_FILTER =
@@ -28,8 +43,13 @@
 val trace = Unsynchronized.ref false
 fun trace_msg msg = if !trace then tracing (msg ()) else ()
 
+(* experimental feature *)
+val term_patterns = false
+
 val respect_no_atp = true
 
+datatype locality = General | Intro | Elim | Simp | Local | Chained
+
 type relevance_override =
   {add: Facts.ref list,
    del: Facts.ref list,
@@ -37,13 +57,22 @@
 
 val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
 
-fun name_thms_pair_from_ref ctxt reserved chained_ths xref =
-  let val ths = ProofContext.get_fact ctxt xref in
-    (fn () => let
-                val name = Facts.string_of_ref xref
-                val name = name |> Symtab.defined reserved name ? quote
-                val chained = forall (member Thm.eq_thm chained_ths) ths
-              in (name, chained) end, ths)
+fun repair_name reserved multi j name =
+  (name |> Symtab.defined reserved name ? quote) ^
+  (if multi then "(" ^ Int.toString j ^ ")" else "")
+
+fun name_thm_pairs_from_ref ctxt reserved chained_ths xref =
+  let
+    val ths = ProofContext.get_fact ctxt xref
+    val name = Facts.string_of_ref xref
+    val multi = length ths > 1
+  in
+    (ths, (1, []))
+    |-> fold (fn th => fn (j, rest) =>
+                 (j + 1, ((repair_name reserved multi j name,
+                          if member Thm.eq_thm chained_ths th then Chained
+                          else General), th) :: rest))
+    |> snd
   end
 
 (***************************************************************)
@@ -52,104 +81,146 @@
 
 (*** constants with types ***)
 
-(*An abstraction of Isabelle types*)
-datatype const_typ =  CTVar | CType of string * const_typ list
+fun order_of_type (Type (@{type_name fun}, [T1, @{typ bool}])) =
+    order_of_type T1 (* cheat: pretend sets are first-order *)
+  | order_of_type (Type (@{type_name fun}, [T1, T2])) =
+    Int.max (order_of_type T1 + 1, order_of_type T2)
+  | order_of_type (Type (_, Ts)) = fold (Integer.max o order_of_type) Ts 0
+  | order_of_type _ = 0
+
+(* An abstraction of Isabelle types and first-order terms *)
+datatype pattern = PVar | PApp of string * pattern list
+datatype ptype = PType of int * pattern list
+
+fun string_for_pattern PVar = "_"
+  | string_for_pattern (PApp (s, ps)) =
+    if null ps then s else s ^ string_for_patterns ps
+and string_for_patterns ps = "(" ^ commas (map string_for_pattern ps) ^ ")"
+fun string_for_ptype (PType (_, ps)) = string_for_patterns ps
 
 (*Is the second type an instance of the first one?*)
-fun match_type (CType(con1,args1)) (CType(con2,args2)) =
-      con1=con2 andalso match_types args1 args2
-  | match_type CTVar _ = true
-  | match_type _ CTVar = false
-and match_types [] [] = true
-  | match_types (a1::as1) (a2::as2) = match_type a1 a2 andalso match_types as1 as2;
+fun match_pattern (PVar, _) = true
+  | match_pattern (PApp _, PVar) = false
+  | match_pattern (PApp (s, ps), PApp (t, qs)) =
+    s = t andalso match_patterns (ps, qs)
+and match_patterns (_, []) = true
+  | match_patterns ([], _) = false
+  | match_patterns (p :: ps, q :: qs) =
+    match_pattern (p, q) andalso match_patterns (ps, qs)
+fun match_ptype (PType (_, ps), PType (_, qs)) = match_patterns (ps, qs)
 
-(*Is there a unifiable constant?*)
-fun const_mem const_tab (c, c_typ) =
-  exists (match_types c_typ) (these (Symtab.lookup const_tab c))
+(* Is there a unifiable constant? *)
+fun pconst_mem f consts (s, ps) =
+  exists (curry (match_ptype o f) ps)
+         (map snd (filter (curry (op =) s o fst) consts))
+fun pconst_hyper_mem f const_tab (s, ps) =
+  exists (curry (match_ptype o f) ps) (these (Symtab.lookup const_tab s))
+
+fun pattern_for_type (Type (s, Ts)) = PApp (s, map pattern_for_type Ts)
+  | pattern_for_type (TFree (s, _)) = PApp (s, [])
+  | pattern_for_type (TVar _) = PVar
 
-(*Maps a "real" type to a const_typ*)
-fun const_typ_of (Type (c,typs)) = CType (c, map const_typ_of typs)
-  | const_typ_of (TFree _) = CTVar
-  | const_typ_of (TVar _) = CTVar
+fun pterm thy t =
+  case strip_comb t of
+    (Const x, ts) => PApp (pconst thy true x ts)
+  | (Free x, ts) => PApp (pconst thy false x ts)
+  | (Var x, []) => PVar
+  | _ => PApp ("?", [])  (* equivalence class of higher-order constructs *)
+(* Pairs a constant with the list of its type instantiations. *)
+and ptype thy const x ts =
+  (if const then map pattern_for_type (these (try (Sign.const_typargs thy) x))
+   else []) @
+  (if term_patterns then map (pterm thy) ts else [])
+and pconst thy const (s, T) ts = (s, ptype thy const (s, T) ts)
+and rich_ptype thy const (s, T) ts =
+  PType (order_of_type T, ptype thy const (s, T) ts)
+and rich_pconst thy const (s, T) ts = (s, rich_ptype thy const (s, T) ts)
 
-(*Pairs a constant with the list of its type instantiations (using const_typ)*)
-fun const_with_typ thy (c,typ) =
-  let val tvars = Sign.const_typargs thy (c,typ) in
-    (c, map const_typ_of tvars) end
-  handle TYPE _ => (c, [])   (*Variable (locale constant): monomorphic*)
+fun string_for_hyper_pconst (s, ps) =
+  s ^ "{" ^ commas (map string_for_ptype ps) ^ "}"
+
+val abs_name = "Sledgehammer.abs"
+val skolem_prefix = "Sledgehammer.sko"
 
-(*Add a const/type pair to the table, but a [] entry means a standard connective,
-  which we ignore.*)
-fun add_const_to_table (c, ctyps) =
-  Symtab.map_default (c, [ctyps])
-                     (fn [] => [] | ctypss => insert (op =) ctyps ctypss)
+(* These are typically simplified away by "Meson.presimplify". Equality is
+   handled specially via "fequal". *)
+val boring_consts =
+  [@{const_name False}, @{const_name True}, @{const_name If}, @{const_name Let},
+   @{const_name HOL.eq}]
+
+(* Add a pconstant to the table, but a [] entry means a standard
+   connective, which we ignore.*)
+fun add_pconst_to_table also_skolem (c, p) =
+  if member (op =) boring_consts c orelse
+     (not also_skolem andalso String.isPrefix skolem_prefix c) then
+    I
+  else
+    Symtab.map_default (c, [p]) (insert (op =) p)
 
 fun is_formula_type T = (T = HOLogic.boolT orelse T = propT)
 
-val fresh_prefix = "Sledgehammer.FRESH."
-val flip = Option.map not
-(* These are typically simplified away by "Meson.presimplify". *)
-val boring_consts =
-  [@{const_name False}, @{const_name True}, @{const_name If}, @{const_name Let}]
-
-fun get_consts thy pos ts =
+fun pconsts_in_terms thy also_skolems pos ts =
   let
+    val flip = Option.map not
     (* We include free variables, as well as constants, to handle locales. For
        each quantifiers that must necessarily be skolemized by the ATP, we
        introduce a fresh constant to simulate the effect of Skolemization. *)
-    fun do_term t =
-      case t of
-        Const x => add_const_to_table (const_with_typ thy x)
-      | Free (s, _) => add_const_to_table (s, [])
-      | t1 $ t2 => fold do_term [t1, t2]
-      | Abs (_, _, t') => do_term t'
-      | _ => I
-    fun do_quantifier will_surely_be_skolemized body_t =
+    fun do_const const (s, T) ts =
+      add_pconst_to_table also_skolems (rich_pconst thy const (s, T) ts)
+      #> fold do_term ts
+    and do_term t =
+      case strip_comb t of
+        (Const x, ts) => do_const true x ts
+      | (Free x, ts) => do_const false x ts
+      | (Abs (_, T, t'), ts) =>
+        (null ts
+         ? add_pconst_to_table true (abs_name, PType (order_of_type T + 1, [])))
+        #> fold do_term (t' :: ts)
+      | (_, ts) => fold do_term ts
+    fun do_quantifier will_surely_be_skolemized abs_T body_t =
       do_formula pos body_t
-      #> (if will_surely_be_skolemized then
-            add_const_to_table (gensym fresh_prefix, [])
+      #> (if also_skolems andalso will_surely_be_skolemized then
+            add_pconst_to_table true
+                         (gensym skolem_prefix, PType (order_of_type abs_T, []))
           else
             I)
     and do_term_or_formula T =
       if is_formula_type T then do_formula NONE else do_term
     and do_formula pos t =
       case t of
-        Const (@{const_name all}, _) $ Abs (_, _, body_t) =>
-        do_quantifier (pos = SOME false) body_t
+        Const (@{const_name all}, _) $ Abs (_, T, t') =>
+        do_quantifier (pos = SOME false) T t'
       | @{const "==>"} $ t1 $ t2 =>
         do_formula (flip pos) t1 #> do_formula pos t2
       | Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 =>
         fold (do_term_or_formula T) [t1, t2]
       | @{const Trueprop} $ t1 => do_formula pos t1
       | @{const Not} $ t1 => do_formula (flip pos) t1
-      | Const (@{const_name All}, _) $ Abs (_, _, body_t) =>
-        do_quantifier (pos = SOME false) body_t
-      | Const (@{const_name Ex}, _) $ Abs (_, _, body_t) =>
-        do_quantifier (pos = SOME true) body_t
-      | @{const "op &"} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
-      | @{const "op |"} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
-      | @{const "op -->"} $ t1 $ t2 =>
+      | Const (@{const_name All}, _) $ Abs (_, T, t') =>
+        do_quantifier (pos = SOME false) T t'
+      | Const (@{const_name Ex}, _) $ Abs (_, T, t') =>
+        do_quantifier (pos = SOME true) T t'
+      | @{const HOL.conj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
+      | @{const HOL.disj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
+      | @{const HOL.implies} $ t1 $ t2 =>
         do_formula (flip pos) t1 #> do_formula pos t2
-      | Const (@{const_name "op ="}, Type (_, [T, _])) $ t1 $ t2 =>
+      | Const (@{const_name HOL.eq}, Type (_, [T, _])) $ t1 $ t2 =>
         fold (do_term_or_formula T) [t1, t2]
       | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])]))
         $ t1 $ t2 $ t3 =>
         do_formula NONE t1 #> fold (do_term_or_formula T) [t2, t3]
-      | Const (@{const_name Ex1}, _) $ Abs (_, _, body_t) =>
-        do_quantifier (is_some pos) body_t
-      | Const (@{const_name Ball}, _) $ t1 $ Abs (_, _, body_t) =>
-        do_quantifier (pos = SOME false)
-                      (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, body_t))
-      | Const (@{const_name Bex}, _) $ t1 $ Abs (_, _, body_t) =>
-        do_quantifier (pos = SOME true)
-                      (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, body_t))
+      | Const (@{const_name Ex1}, _) $ Abs (_, T, t') =>
+        do_quantifier (is_some pos) T t'
+      | Const (@{const_name Ball}, _) $ t1 $ Abs (_, T, t') =>
+        do_quantifier (pos = SOME false) T
+                      (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, t'))
+      | Const (@{const_name Bex}, _) $ t1 $ Abs (_, T, t') =>
+        do_quantifier (pos = SOME true) T
+                      (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, t'))
       | (t0 as Const (_, @{typ bool})) $ t1 =>
         do_term t0 #> do_formula pos t1  (* theory constant *)
       | _ => do_term t
-  in
-    Symtab.empty |> fold (Symtab.update o rpair []) boring_consts
-                 |> fold (do_formula pos) ts
-  end
+  in Symtab.empty |> fold (do_formula pos) ts end
 
 (*Inserts a dummy "constant" referring to the theory name, so that relevance
   takes the given theory into account.*)
@@ -164,235 +235,292 @@
 
 (**** Constant / Type Frequencies ****)
 
-(*A two-dimensional symbol table counts frequencies of constants. It's keyed first by
-  constant name and second by its list of type instantiations. For the latter, we need
-  a linear ordering on type const_typ list.*)
-
-local
+(* A two-dimensional symbol table counts frequencies of constants. It's keyed
+   first by constant name and second by its list of type instantiations. For the
+   latter, we need a linear ordering on "pattern list". *)
 
-fun cons_nr CTVar = 0
-  | cons_nr (CType _) = 1;
-
-in
-
-fun const_typ_ord TU =
-  case TU of
-    (CType (a, Ts), CType (b, Us)) =>
-      (case fast_string_ord(a,b) of EQUAL => dict_ord const_typ_ord (Ts,Us) | ord => ord)
-  | (T, U) => int_ord (cons_nr T, cons_nr U);
+fun pattern_ord p =
+  case p of
+    (PVar, PVar) => EQUAL
+  | (PVar, PApp _) => LESS
+  | (PApp _, PVar) => GREATER
+  | (PApp q1, PApp q2) =>
+    prod_ord fast_string_ord (dict_ord pattern_ord) (q1, q2)
+fun ptype_ord (PType p, PType q) =
+  prod_ord (dict_ord pattern_ord) int_ord (swap p, swap q)
 
-end;
+structure PType_Tab = Table(type key = ptype val ord = ptype_ord)
 
-structure CTtab = Table(type key = const_typ list val ord = dict_ord const_typ_ord);
-
-fun count_axiom_consts theory_relevant thy (_, th) =
+fun count_axiom_consts theory_relevant thy =
   let
-    fun do_const (a, T) =
-      let val (c, cts) = const_with_typ thy (a, T) in
-        (* Two-dimensional table update. Constant maps to types maps to
-           count. *)
-        CTtab.map_default (cts, 0) (Integer.add 1)
-        |> Symtab.map_default (c, CTtab.empty)
-      end
-    fun do_term (Const x) = do_const x
-      | do_term (Free x) = do_const x
-      | do_term (t $ u) = do_term t #> do_term u
-      | do_term (Abs (_, _, t)) = do_term t
-      | do_term _ = I
-  in th |> theory_const_prop_of theory_relevant |> do_term end
+    fun do_const const (s, T) ts =
+      (* Two-dimensional table update. Constant maps to types maps to count. *)
+      PType_Tab.map_default (rich_ptype thy const (s, T) ts, 0) (Integer.add 1)
+      |> Symtab.map_default (s, PType_Tab.empty)
+      #> fold do_term ts
+    and do_term t =
+      case strip_comb t of
+        (Const x, ts) => do_const true x ts
+      | (Free x, ts) => do_const false x ts
+      | (Abs (_, _, t'), ts) => fold do_term (t' :: ts)
+      | (_, ts) => fold do_term ts
+  in do_term o theory_const_prop_of theory_relevant o snd end
 
 
 (**** Actual Filtering Code ****)
 
+fun pow_int x 0 = 1.0
+  | pow_int x 1 = x
+  | pow_int x n = if n > 0 then x * pow_int x (n - 1) else pow_int x (n + 1) / x
+
 (*The frequency of a constant is the sum of those of all instances of its type.*)
-fun const_frequency const_tab (c, cts) =
-  CTtab.fold (fn (cts', m) => match_types cts cts' ? Integer.add m)
-             (the (Symtab.lookup const_tab c)) 0
-  handle Option.Option => 0
+fun pconst_freq match const_tab (c, ps) =
+  PType_Tab.fold (fn (qs, m) => match (ps, qs) ? Integer.add m)
+                 (the (Symtab.lookup const_tab c)) 0
 
 
 (* A surprising number of theorems contain only a few significant constants.
    These include all induction rules, and other general theorems. *)
 
 (* "log" seems best in practice. A constant function of one ignores the constant
-   frequencies. *)
-fun rel_log (x : real) = 1.0 + 2.0 / Math.ln (x + 1.0)
-fun irrel_log (x : real) = Math.ln (x + 19.0) / 6.4
+   frequencies. Rare constants give more points if they are relevant than less
+   rare ones. *)
+fun rel_weight_for order freq = 1.0 + 2.0 / Math.ln (Real.fromInt freq + 1.0)
+
+(* FUDGE *)
+val worse_irrel_freq = Unsynchronized.ref 100.0
+val higher_order_irrel_weight = Unsynchronized.ref 1.05
+
+(* Irrelevant constants are treated differently. We associate lower penalties to
+   very rare constants and very common ones -- the former because they can't
+   lead to the inclusion of too many new facts, and the latter because they are
+   so common as to be of little interest. *)
+fun irrel_weight_for order freq =
+  let val (k, x) = !worse_irrel_freq |> `Real.ceil in
+    (if freq < k then Math.ln (Real.fromInt (freq + 1)) / Math.ln x
+     else rel_weight_for order freq / rel_weight_for order k)
+    * pow_int (!higher_order_irrel_weight) (order - 1)
+  end
+
+(* FUDGE *)
+val abs_rel_weight = Unsynchronized.ref 0.5
+val abs_irrel_weight = Unsynchronized.ref 2.0
+val skolem_irrel_weight = Unsynchronized.ref 0.75
 
 (* Computes a constant's weight, as determined by its frequency. *)
-val rel_const_weight = rel_log o real oo const_frequency
-val irrel_const_weight = irrel_log o real oo const_frequency
-(* fun irrel_const_weight _ _ = 1.0  FIXME: OLD CODE *)
+fun generic_pconst_weight abs_weight skolem_weight weight_for f const_tab
+                          (c as (s, PType (m, _))) =
+  if s = abs_name then abs_weight
+  else if String.isPrefix skolem_prefix s then skolem_weight
+  else weight_for m (pconst_freq (match_ptype o f) const_tab c)
+
+fun rel_pconst_weight const_tab =
+  generic_pconst_weight (!abs_rel_weight) 0.0 rel_weight_for I const_tab
+fun irrel_pconst_weight const_tab =
+  generic_pconst_weight (!abs_irrel_weight) (!skolem_irrel_weight)
+                        irrel_weight_for swap const_tab
+
+(* FUDGE *)
+val intro_bonus = Unsynchronized.ref 0.15
+val elim_bonus = Unsynchronized.ref 0.15
+val simp_bonus = Unsynchronized.ref 0.15
+val local_bonus = Unsynchronized.ref 0.55
+val chained_bonus = Unsynchronized.ref 1.5
+
+fun locality_bonus General = 0.0
+  | locality_bonus Intro = !intro_bonus
+  | locality_bonus Elim = !elim_bonus
+  | locality_bonus Simp = !simp_bonus
+  | locality_bonus Local = !local_bonus
+  | locality_bonus Chained = !chained_bonus
 
-fun axiom_weight const_tab relevant_consts axiom_consts =
-  let
-    val (rel, irrel) = List.partition (const_mem relevant_consts) axiom_consts
-    val rel_weight = fold (curry Real.+ o rel_const_weight const_tab) rel 0.0
-    val irrel_weight = fold (curry Real.+ o irrel_const_weight const_tab) irrel 0.0
-    val res = rel_weight / (rel_weight + irrel_weight)
-  in if Real.isFinite res then res else 0.0 end
+fun axiom_weight loc const_tab relevant_consts axiom_consts =
+  case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts)
+                    ||> filter_out (pconst_hyper_mem swap relevant_consts) of
+    ([], _) => 0.0
+  | (rel, irrel) =>
+    let
+      val irrel = irrel |> filter_out (pconst_mem swap rel)
+      val rel_weight =
+        0.0 |> fold (curry (op +) o rel_pconst_weight const_tab) rel
+      val irrel_weight =
+        ~ (locality_bonus loc)
+        |> fold (curry (op +) o irrel_pconst_weight const_tab) irrel
+      val res = rel_weight / (rel_weight + irrel_weight)
+    in if Real.isFinite res then res else 0.0 end
 
-(* OLD CODE:
-(*Relevant constants are weighted according to frequency,
-  but irrelevant constants are simply counted. Otherwise, Skolem functions,
-  which are rare, would harm a formula's chances of being picked.*)
-fun axiom_weight const_tab relevant_consts axiom_consts =
-  let
-    val rel = filter (const_mem relevant_consts) axiom_consts
-    val rel_weight = fold (curry Real.+ o rel_const_weight const_tab) rel 0.0
-    val res = rel_weight / (rel_weight + real (length axiom_consts - length rel))
-  in if Real.isFinite res then res else 0.0 end
+(* FIXME: experiment
+fun debug_axiom_weight loc const_tab relevant_consts axiom_consts =
+  case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts)
+                    ||> filter_out (pconst_hyper_mem swap relevant_consts) of
+    ([], _) => 0.0
+  | (rel, irrel) =>
+    let
+      val irrel = irrel |> filter_out (pconst_mem swap rel)
+      val rels_weight =
+        0.0 |> fold (curry (op +) o rel_pconst_weight const_tab) rel
+      val irrels_weight =
+        ~ (locality_bonus loc)
+        |> fold (curry (op +) o irrel_pconst_weight const_tab) irrel
+val _ = tracing (PolyML.makestring ("REL: ", map (`(rel_pconst_weight const_tab)) rel))
+val _ = tracing (PolyML.makestring ("IRREL: ", map (`(irrel_pconst_weight const_tab)) irrel))
+      val res = rels_weight / (rels_weight + irrels_weight)
+    in if Real.isFinite res then res else 0.0 end
 *)
 
-(*Multiplies out to a list of pairs: 'a * 'b list -> ('a * 'b) list -> ('a * 'b) list*)
-fun add_expand_pairs (x, ys) xys = List.foldl (fn (y,acc) => (x,y)::acc) xys ys
-
-fun consts_of_term thy t =
-  Symtab.fold add_expand_pairs (get_consts thy (SOME true) [t]) []
-
+fun pconsts_in_axiom thy t =
+  Symtab.fold (fn (s, pss) => fold (cons o pair s) pss)
+              (pconsts_in_terms thy true (SOME true) [t]) []
 fun pair_consts_axiom theory_relevant thy axiom =
-  (axiom, axiom |> snd |> theory_const_prop_of theory_relevant
-                |> consts_of_term thy)
-
-exception CONST_OR_FREE of unit
-
-fun dest_Const_or_Free (Const x) = x
-  | dest_Const_or_Free (Free x) = x
-  | dest_Const_or_Free _ = raise CONST_OR_FREE ()
-
-(*Look for definitions of the form f ?x1 ... ?xn = t, but not reversed.*)
-fun defines thy thm gctypes =
-    let val tm = prop_of thm
-        fun defs lhs rhs =
-            let val (rator,args) = strip_comb lhs
-                val ct = const_with_typ thy (dest_Const_or_Free rator)
-            in
-              forall is_Var args andalso const_mem gctypes ct andalso
-              subset (op =) (Term.add_vars rhs [], Term.add_vars lhs [])
-            end
-            handle CONST_OR_FREE () => false
-    in
-        case tm of
-          @{const Trueprop} $ (Const (@{const_name "op ="}, _) $ lhs $ rhs) =>
-            defs lhs rhs
-        | _ => false
-    end;
+  case axiom |> snd |> theory_const_prop_of theory_relevant
+             |> pconsts_in_axiom thy of
+    [] => NONE
+  | consts => SOME ((axiom, consts), NONE)
 
 type annotated_thm =
-  ((unit -> string * bool) * thm) * (string * const_typ list) list
+  (((unit -> string) * locality) * thm) * (string * ptype) list
 
-(*For a reverse sort, putting the largest values first.*)
-fun compare_pairs ((_, w1), (_, w2)) = Real.compare (w2, w1)
+(* FUDGE *)
+val max_imperfect = Unsynchronized.ref 11.5
+val max_imperfect_exp = Unsynchronized.ref 1.0
 
-(* Limit the number of new facts, to prevent runaway acceptance. *)
-fun take_best max_new (new_pairs : (annotated_thm * real) list) =
-  let val nnew = length new_pairs in
-    if nnew <= max_new then
-      (map #1 new_pairs, [])
-    else
-      let
-        val new_pairs = sort compare_pairs new_pairs
-        val accepted = List.take (new_pairs, max_new)
-      in
-        trace_msg (fn () => ("Number of candidates, " ^ Int.toString nnew ^
-                       ", exceeds the limit of " ^ Int.toString max_new));
-        trace_msg (fn () => ("Effective pass mark: " ^ Real.toString (#2 (List.last accepted))));
-        trace_msg (fn () => "Actually passed: " ^
-          space_implode ", " (map (fst o (fn f => f ()) o fst o fst o fst) accepted));
-        (map #1 accepted, List.drop (new_pairs, max_new))
-      end
-  end;
+fun take_most_relevant max_relevant remaining_max
+                       (candidates : (annotated_thm * real) list) =
+  let
+    val max_imperfect =
+      Real.ceil (Math.pow (!max_imperfect,
+                    Math.pow (Real.fromInt remaining_max
+                              / Real.fromInt max_relevant, !max_imperfect_exp)))
+    val (perfect, imperfect) =
+      candidates |> sort (Real.compare o swap o pairself snd)
+                 |> take_prefix (fn (_, w) => w > 0.99999)
+    val ((accepts, more_rejects), rejects) =
+      chop max_imperfect imperfect |>> append perfect |>> chop remaining_max
+  in
+    trace_msg (fn () =>
+        "Actually passed (" ^ Int.toString (length accepts) ^ " of " ^
+        Int.toString (length candidates) ^ "): " ^
+        (accepts |> map (fn ((((name, _), _), _), weight) =>
+                            name () ^ " [" ^ Real.toString weight ^ "]")
+                 |> commas));
+    (accepts, more_rejects @ rejects)
+  end
 
-val threshold_divisor = 2.0
-val ridiculous_threshold = 0.1
+fun if_empty_replace_with_locality thy axioms loc tab =
+  if Symtab.is_empty tab then
+    pconsts_in_terms thy false (SOME false)
+        (map_filter (fn ((_, loc'), th) =>
+                        if loc' = loc then SOME (prop_of th) else NONE) axioms)
+  else
+    tab
+
+(* FUDGE *)
+val threshold_divisor = Unsynchronized.ref 2.0
+val ridiculous_threshold = Unsynchronized.ref 0.1
 
-fun relevance_filter ctxt relevance_threshold relevance_convergence
-                     defs_relevant max_new theory_relevant
+fun relevance_filter ctxt threshold0 decay max_relevant theory_relevant
                      ({add, del, ...} : relevance_override) axioms goal_ts =
-  if relevance_threshold > 1.0 then
-    []
-  else if relevance_threshold < 0.0 then
-    axioms
-  else
-    let
-      val thy = ProofContext.theory_of ctxt
-      val const_tab = fold (count_axiom_consts theory_relevant thy) axioms
-                           Symtab.empty
-      val goal_const_tab = get_consts thy (SOME false) goal_ts
-      val _ =
-        trace_msg (fn () => "Initial constants: " ^
-                            commas (goal_const_tab
-                                    |> Symtab.dest
-                                    |> filter (curry (op <>) [] o snd)
-                                    |> map fst))
-      val add_thms = maps (ProofContext.get_fact ctxt) add
-      val del_thms = maps (ProofContext.get_fact ctxt) del
-      fun iter j threshold rel_const_tab =
-        let
-          fun relevant ([], rejects) [] =
-              (* Nothing was added this iteration. *)
-              if j = 0 andalso threshold >= ridiculous_threshold then
-                (* First iteration? Try again. *)
-                iter 0 (threshold / threshold_divisor) rel_const_tab
-                     (map (apsnd SOME) rejects)
+  let
+    val thy = ProofContext.theory_of ctxt
+    val const_tab =
+      fold (count_axiom_consts theory_relevant thy) axioms Symtab.empty
+    val goal_const_tab =
+      pconsts_in_terms thy false (SOME false) goal_ts
+      |> fold (if_empty_replace_with_locality thy axioms) [Chained, Local]
+    val add_thms = maps (ProofContext.get_fact ctxt) add
+    val del_thms = maps (ProofContext.get_fact ctxt) del
+    fun iter j remaining_max threshold rel_const_tab hopeless hopeful =
+      let
+        fun game_over rejects =
+          (* Add "add:" facts. *)
+          if null add_thms then
+            []
+          else
+            map_filter (fn ((p as (_, th), _), _) =>
+                           if member Thm.eq_thm add_thms th then SOME p
+                           else NONE) rejects
+        fun relevant [] rejects [] =
+            (* Nothing has been added this iteration. *)
+            if j = 0 andalso threshold >= !ridiculous_threshold then
+              (* First iteration? Try again. *)
+              iter 0 max_relevant (threshold / !threshold_divisor) rel_const_tab
+                   hopeless hopeful
+            else
+              game_over (rejects @ hopeless)
+          | relevant candidates rejects [] =
+            let
+              val (accepts, more_rejects) =
+                take_most_relevant max_relevant remaining_max candidates
+              val rel_const_tab' =
+                rel_const_tab
+                |> fold (add_pconst_to_table false) (maps (snd o fst) accepts)
+              fun is_dirty (c, _) =
+                Symtab.lookup rel_const_tab' c <> Symtab.lookup rel_const_tab c
+              val (hopeful_rejects, hopeless_rejects) =
+                 (rejects @ hopeless, ([], []))
+                 |-> fold (fn (ax as (_, consts), old_weight) =>
+                              if exists is_dirty consts then
+                                apfst (cons (ax, NONE))
+                              else
+                                apsnd (cons (ax, old_weight)))
+                 |>> append (more_rejects
+                             |> map (fn (ax as (_, consts), old_weight) =>
+                                        (ax, if exists is_dirty consts then NONE
+                                             else SOME old_weight)))
+              val threshold =
+                1.0 - (1.0 - threshold)
+                      * Math.pow (decay, Real.fromInt (length accepts))
+              val remaining_max = remaining_max - length accepts
+            in
+              trace_msg (fn () => "New or updated constants: " ^
+                  commas (rel_const_tab' |> Symtab.dest
+                          |> subtract (op =) (rel_const_tab |> Symtab.dest)
+                          |> map string_for_hyper_pconst));
+              map (fst o fst) accepts @
+              (if remaining_max = 0 then
+                 game_over (hopeful_rejects @ map (apsnd SOME) hopeless_rejects)
+               else
+                 iter (j + 1) remaining_max threshold rel_const_tab'
+                      hopeless_rejects hopeful_rejects)
+            end
+          | relevant candidates rejects
+                     (((ax as (((_, loc), th), axiom_consts)), cached_weight)
+                      :: hopeful) =
+            let
+              val weight =
+                case cached_weight of
+                  SOME w => w
+                | NONE => axiom_weight loc const_tab rel_const_tab axiom_consts
+(* FIXME: experiment
+val name = fst (fst (fst ax)) ()
+val _ = if String.isSubstring "positive_minus" name orelse String.isSubstring "not_exp_le_zero" name then
+tracing ("*** " ^ name ^ PolyML.makestring (debug_axiom_weight loc const_tab rel_const_tab axiom_consts))
+else
+()
+*)
+            in
+              if weight >= threshold then
+                relevant ((ax, weight) :: candidates) rejects hopeful
               else
-                (* Add "add:" facts. *)
-                if null add_thms then
-                  []
-                else
-                  map_filter (fn ((p as (_, th), _), _) =>
-                                 if member Thm.eq_thm add_thms th then SOME p
-                                 else NONE) rejects
-            | relevant (new_pairs, rejects) [] =
-              let
-                val (new_rels, more_rejects) = take_best max_new new_pairs
-                val rel_const_tab' =
-                  rel_const_tab |> fold add_const_to_table (maps snd new_rels)
-                fun is_dirty c =
-                  const_mem rel_const_tab' c andalso
-                  not (const_mem rel_const_tab c)
-                val rejects =
-                  more_rejects @ rejects
-                  |> map (fn (ax as (_, consts), old_weight) =>
-                             (ax, if exists is_dirty consts then NONE
-                                  else SOME old_weight))
-                val threshold =
-                  threshold + (1.0 - threshold) * relevance_convergence
-              in
-                trace_msg (fn () => "relevant this iteration: " ^
-                                    Int.toString (length new_rels));
-                map #1 new_rels @ iter (j + 1) threshold rel_const_tab' rejects
-              end
-            | relevant (new_rels, rejects)
-                       (((ax as ((name, th), axiom_consts)), cached_weight)
-                        :: rest) =
-              let
-                val weight =
-                  case cached_weight of
-                    SOME w => w
-                  | NONE => axiom_weight const_tab rel_const_tab axiom_consts
-              in
-                if weight >= threshold orelse
-                   (defs_relevant andalso defines thy th rel_const_tab) then
-                  (trace_msg (fn () =>
-                       fst (name ()) ^ " passes: " ^ Real.toString weight
-                       ^ " consts: " ^ commas (map fst axiom_consts));
-                   relevant ((ax, weight) :: new_rels, rejects) rest)
-                else
-                  relevant (new_rels, (ax, weight) :: rejects) rest
-              end
-          in
-            trace_msg (fn () => "relevant_facts, current threshold: " ^
-                                Real.toString threshold);
-            relevant ([], [])
-          end
-    in
-      axioms |> filter_out (member Thm.eq_thm del_thms o snd)
-             |> map (rpair NONE o pair_consts_axiom theory_relevant thy)
-             |> iter 0 relevance_threshold goal_const_tab
-             |> tap (fn res => trace_msg (fn () =>
+                relevant candidates ((ax, weight) :: rejects) hopeful
+            end
+        in
+          trace_msg (fn () =>
+              "ITERATION " ^ string_of_int j ^ ": current threshold: " ^
+              Real.toString threshold ^ ", constants: " ^
+              commas (rel_const_tab |> Symtab.dest
+                      |> filter (curry (op <>) [] o snd)
+                      |> map string_for_hyper_pconst));
+          relevant [] [] hopeful
+        end
+  in
+    axioms |> filter_out (member Thm.eq_thm del_thms o snd)
+           |> map_filter (pair_consts_axiom theory_relevant thy)
+           |> iter 0 max_relevant threshold0 goal_const_tab []
+           |> tap (fn res => trace_msg (fn () =>
                                 "Total relevant: " ^ Int.toString (length res)))
-    end
+  end
+
 
 (***************************************************************)
 (* Retrieving and filtering lemmas                             *)
@@ -410,9 +538,9 @@
      String.isSuffix "_def" a  orelse  String.isSuffix "_defs" a
   end;
 
-fun make_fact_table xs =
-  fold (Termtab.update o `(prop_of o snd)) xs Termtab.empty
-fun make_unique xs = Termtab.fold (cons o snd) (make_fact_table xs) []
+fun mk_fact_table f xs =
+  fold (Termtab.update o `(prop_of o f)) xs Termtab.empty
+fun uniquify xs = Termtab.fold (cons o snd) (mk_fact_table snd xs) []
 
 (* FIXME: put other record thms here, or declare as "no_atp" *)
 val multi_base_blacklist =
@@ -453,12 +581,22 @@
 val exists_sledgehammer_const =
   exists_Const (fn (s, _) => String.isPrefix sledgehammer_prefix s)
 
-fun is_strange_theorem th =
+(* FIXME: make more reliable *)
+val exists_low_level_class_const =
+  exists_Const (fn (s, _) =>
+     String.isSubstring (Long_Name.separator ^ "class" ^ Long_Name.separator) s)
+
+fun is_metastrange_theorem th =
   case head_of (concl_of th) of
       Const (a, _) => (a <> @{const_name Trueprop} andalso
                        a <> @{const_name "=="})
     | _ => false
 
+fun is_that_fact th =
+  String.isSuffix (Long_Name.separator ^ Obtain.thatN) (Thm.get_name_hint th)
+  andalso exists_subterm (fn Free (s, _) => s = Name.skolem Auto_Bind.thesisN
+                           | _ => false) (prop_of th)
+
 val type_has_top_sort =
   exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false)
 
@@ -491,13 +629,13 @@
       | (_, @{const "==>"} $ t1 $ t2) =>
         do_formula (not pos) t1 andalso
         (t2 = @{prop False} orelse do_formula pos t2)
-      | (_, @{const "op -->"} $ t1 $ t2) =>
+      | (_, @{const HOL.implies} $ t1 $ t2) =>
         do_formula (not pos) t1 andalso
         (t2 = @{const False} orelse do_formula pos t2)
       | (_, @{const Not} $ t1) => do_formula (not pos) t1
-      | (true, @{const "op |"} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
-      | (false, @{const "op &"} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
-      | (true, Const (@{const_name "op ="}, _) $ t1 $ t2) => do_equals t1 t2
+      | (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
+      | (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
+      | (true, Const (@{const_name HOL.eq}, _) $ t1 $ t2) => do_equals t1 t2
       | (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
       | _ => false
   in do_formula true end
@@ -528,27 +666,42 @@
   let val t = prop_of thm in
     is_formula_too_complex t orelse exists_type type_has_top_sort t orelse
     is_dangerous_term full_types t orelse exists_sledgehammer_const t orelse
-    is_strange_theorem thm
+    exists_low_level_class_const t orelse is_metastrange_theorem thm orelse
+    is_that_fact thm
   end
 
+fun clasimpset_rules_of ctxt =
+  let
+    val {safeIs, safeEs, hazIs, hazEs, ...} = ctxt |> claset_of |> rep_cs
+    val intros = safeIs @ hazIs
+    val elims = map Classical.classical_rule (safeEs @ hazEs)
+    val simps = ctxt |> simpset_of |> dest_ss |> #simps |> map snd
+  in (mk_fact_table I intros, mk_fact_table I elims, mk_fact_table I simps) end
+
 fun all_name_thms_pairs ctxt reserved full_types add_thms chained_ths =
   let
-    val is_chained = member Thm.eq_thm chained_ths
-    val global_facts = PureThy.facts_of (ProofContext.theory_of ctxt)
+    val thy = ProofContext.theory_of ctxt
+    val global_facts = PureThy.facts_of thy
     val local_facts = ProofContext.facts_of ctxt
     val named_locals = local_facts |> Facts.dest_static []
-    (* Unnamed, not chained formulas with schematic variables are omitted,
-       because they are rejected by the backticks (`...`) parser for some
-       reason. *)
-    fun is_bad_unnamed_local th =
-      exists (fn (_, ths) => member Thm.eq_thm ths th) named_locals orelse
-      (exists_subterm is_Var (prop_of th) andalso not (is_chained th))
+    val is_chained = member Thm.eq_thm chained_ths
+    val (intros, elims, simps) =
+      if exists (curry (op <) 0.0) [!intro_bonus, !elim_bonus, !simp_bonus] then
+        clasimpset_rules_of ctxt
+      else
+        (Termtab.empty, Termtab.empty, Termtab.empty)
+    (* Unnamed nonchained formulas with schematic variables are omitted, because
+       they are rejected by the backticks (`...`) parser for some reason. *)
+    fun is_good_unnamed_local th =
+      not (Thm.has_name_hint th) andalso
+      (not (exists_subterm is_Var (prop_of th)) orelse (is_chained th)) andalso
+      forall (fn (_, ths) => not (member Thm.eq_thm ths th)) named_locals
     val unnamed_locals =
-      local_facts |> Facts.props |> filter_out is_bad_unnamed_local
-                  |> map (pair "" o single)
+      union Thm.eq_thm (Facts.props local_facts) chained_ths
+      |> filter is_good_unnamed_local |> map (pair "" o single)
     val full_space =
-      Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts);
-    fun add_valid_facts foldx facts =
+      Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts)
+    fun add_facts global foldx facts =
       foldx (fn (name0, ths) =>
         if name0 <> "" andalso
            forall (not o member Thm.eq_thm add_thms) ths andalso
@@ -563,6 +716,8 @@
             fun backquotify th =
               "`" ^ Print_Mode.setmp [Print_Mode.input]
                                  (Syntax.string_of_term ctxt) (prop_of th) ^ "`"
+              |> String.translate (fn c => if Char.isPrint c then str c else "")
+              |> simplify_spaces
             fun check_thms a =
               case try (ProofContext.get_thms ctxt) a of
                 NONE => false
@@ -575,66 +730,71 @@
                      not (member Thm.eq_thm add_thms th) then
                     rest
                   else
-                    (fn () =>
-                        (if name0 = "" then
-                           th |> backquotify
-                         else
-                           let
-                             val name1 = Facts.extern facts name0
-                             val name2 = Name_Space.extern full_space name0
-                           in
-                             case find_first check_thms [name1, name2, name0] of
-                               SOME name =>
-                               let
-                                 val name =
-                                   name |> Symtab.defined reserved name ? quote
-                               in
-                                 if multi then name ^ "(" ^ Int.toString j ^ ")"
-                                 else name
-                               end
-                             | NONE => ""
-                           end, is_chained th), (multi, th)) :: rest)) ths
+                    (((fn () =>
+                          if name0 = "" then
+                            th |> backquotify
+                          else
+                            let
+                              val name1 = Facts.extern facts name0
+                              val name2 = Name_Space.extern full_space name0
+                            in
+                              case find_first check_thms [name1, name2, name0] of
+                                SOME name => repair_name reserved multi j name
+                              | NONE => ""
+                            end),
+                      let val t = prop_of th in
+                        if is_chained th then Chained
+                        else if not global then Local
+                        else if Termtab.defined intros t then Intro
+                        else if Termtab.defined elims t then Elim
+                        else if Termtab.defined simps t then Simp
+                        else General
+                      end),
+                      (multi, th)) :: rest)) ths
             #> snd
           end)
   in
-    [] |> add_valid_facts fold local_facts (unnamed_locals @ named_locals)
-       |> add_valid_facts Facts.fold_static global_facts global_facts
+    [] |> add_facts false fold local_facts (unnamed_locals @ named_locals)
+       |> add_facts true Facts.fold_static global_facts global_facts
   end
 
 (* The single-name theorems go after the multiple-name ones, so that single
    names are preferred when both are available. *)
 fun name_thm_pairs ctxt respect_no_atp =
-  List.partition (fst o snd) #> op @
-  #> map (apsnd snd)
+  List.partition (fst o snd) #> op @ #> map (apsnd snd)
   #> respect_no_atp ? filter_out (No_ATPs.member ctxt o snd)
 
 (***************************************************************)
 (* ATP invocation methods setup                                *)
 (***************************************************************)
 
-fun relevant_facts full_types relevance_threshold relevance_convergence
-                   defs_relevant max_new theory_relevant
-                   (relevance_override as {add, del, only})
-                   (ctxt, (chained_ths, _)) hyp_ts concl_t =
+fun relevant_facts ctxt full_types (threshold0, threshold1) max_relevant
+                   theory_relevant (relevance_override as {add, del, only})
+                   chained_ths hyp_ts concl_t =
   let
+    val decay = Math.pow ((1.0 - threshold1) / (1.0 - threshold0),
+                          1.0 / Real.fromInt (max_relevant + 1))
     val add_thms = maps (ProofContext.get_fact ctxt) add
     val reserved = reserved_isar_keyword_table ()
     val axioms =
       (if only then
-         maps ((fn (n, ths) => map (pair n o pair false) ths)
-               o name_thms_pair_from_ref ctxt reserved chained_ths) add
+         maps (map (fn ((name, loc), th) => ((K name, loc), (true, th)))
+               o name_thm_pairs_from_ref ctxt reserved chained_ths) add
        else
          all_name_thms_pairs ctxt reserved full_types add_thms chained_ths)
       |> name_thm_pairs ctxt (respect_no_atp andalso not only)
-      |> make_unique
+      |> uniquify
   in
     trace_msg (fn () => "Considering " ^ Int.toString (length axioms) ^
                         " theorems");
-    relevance_filter ctxt relevance_threshold relevance_convergence
-                     defs_relevant max_new theory_relevant relevance_override
-                     axioms (concl_t :: hyp_ts)
-    |> map (apfst (fn f => f ()))
-    |> sort_wrt (fst o fst)
+    (if threshold0 > 1.0 orelse threshold0 > threshold1 then
+       []
+     else if threshold0 < 0.0 then
+       axioms
+     else
+       relevance_filter ctxt threshold0 decay max_relevant theory_relevant
+                        relevance_override axioms (concl_t :: hyp_ts))
+    |> map (apfst (apfst (fn f => f ())))
   end
 
 end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimize.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimize.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -7,11 +7,12 @@
 
 signature SLEDGEHAMMER_FACT_MINIMIZE =
 sig
+  type locality = Sledgehammer_Fact_Filter.locality
   type params = Sledgehammer.params
 
   val minimize_theorems :
-    params -> int -> int -> Proof.state -> ((string * bool) * thm list) list
-    -> ((string * bool) * thm list) list option * string
+    params -> int -> int -> Proof.state -> ((string * locality) * thm list) list
+    -> ((string * locality) * thm list) list option * string
   val run_minimize : params -> int -> Facts.ref list -> Proof.state -> unit
 end;
 
@@ -40,24 +41,20 @@
        "")
   end
 
-fun test_theorems ({debug, verbose, overlord, atps, full_types,
-                    relevance_threshold, relevance_convergence,
-                    defs_relevant, isar_proof, isar_shrink_factor, ...}
-                   : params)
+fun test_theorems ({debug, verbose, overlord, atps, full_types, isar_proof,
+                    isar_shrink_factor, ...} : params)
                   (prover : prover) explicit_apply timeout subgoal state
-                  name_thms_pairs =
+                  axioms =
   let
     val _ =
-      priority ("Testing " ^ n_theorems (map fst name_thms_pairs) ^ "...")
+      priority ("Testing " ^ n_theorems (map fst axioms) ^ "...")
     val params =
       {debug = debug, verbose = verbose, overlord = overlord, atps = atps,
        full_types = full_types, explicit_apply = explicit_apply,
-       relevance_threshold = relevance_threshold,
-       relevance_convergence = relevance_convergence,
-       max_relevant_per_iter = NONE, theory_relevant = NONE,
-       defs_relevant = defs_relevant, isar_proof = isar_proof,
+       relevance_thresholds = (1.01, 1.01), max_relevant = NONE,
+       theory_relevant = NONE, isar_proof = isar_proof,
        isar_shrink_factor = isar_shrink_factor, timeout = timeout}
-    val axioms = maps (fn (n, ths) => map (pair n) ths) name_thms_pairs
+    val axioms = maps (fn (n, ths) => map (pair n) ths) axioms
     val {context = ctxt, facts, goal} = Proof.goal state
     val problem =
      {subgoal = subgoal, goal = (ctxt, (facts, goal)),
@@ -67,7 +64,7 @@
   in
     priority (case outcome of
                 NONE =>
-                if length used_thm_names = length name_thms_pairs then
+                if length used_thm_names = length axioms then
                   "Found proof."
                 else
                   "Found proof with " ^ n_theorems used_thm_names ^ "."
@@ -93,10 +90,9 @@
 val fudge_msecs = 1000
 
 fun minimize_theorems {atps = [], ...} _ _ _ _ = error "No ATP is set."
-  | minimize_theorems
-        (params as {debug, atps = atp :: _, full_types, isar_proof,
-                    isar_shrink_factor, timeout, ...})
-        i n state name_thms_pairs =
+  | minimize_theorems (params as {debug, atps = atp :: _, full_types,
+                                  isar_proof, isar_shrink_factor, timeout, ...})
+                      i n state axioms =
   let
     val thy = Proof.theory_of state
     val prover = get_prover_fun thy atp
@@ -106,13 +102,12 @@
     val (_, hyp_ts, concl_t) = strip_subgoal goal i
     val explicit_apply =
       not (forall (Meson.is_fol_term thy)
-                  (concl_t :: hyp_ts @
-                   maps (map prop_of o snd) name_thms_pairs))
+                  (concl_t :: hyp_ts @ maps (map prop_of o snd) axioms))
     fun do_test timeout =
       test_theorems params prover explicit_apply timeout i state
     val timer = Timer.startRealTimer ()
   in
-    (case do_test timeout name_thms_pairs of
+    (case do_test timeout axioms of
        result as {outcome = NONE, pool, used_thm_names,
                   conjecture_shape, ...} =>
        let
@@ -122,11 +117,11 @@
            |> Time.fromMilliseconds
          val (min_thms, {proof, axiom_names, ...}) =
            sublinear_minimize (do_test new_timeout)
-               (filter_used_facts used_thm_names name_thms_pairs) ([], result)
+               (filter_used_facts used_thm_names axioms) ([], result)
          val n = length min_thms
          val _ = priority (cat_lines
            ["Minimized: " ^ string_of_int n ^ " theorem" ^ plural_s n] ^
-            (case length (filter (snd o fst) min_thms) of
+            (case length (filter (curry (op =) Chained o snd o fst) min_thms) of
                0 => ""
              | n => " (including " ^ Int.toString n ^ " chained)") ^ ".")
        in
@@ -154,15 +149,14 @@
     val ctxt = Proof.context_of state
     val reserved = reserved_isar_keyword_table ()
     val chained_ths = #facts (Proof.goal state)
-    val name_thms_pairs =
-      map (apfst (fn f => f ())
-           o name_thms_pair_from_ref ctxt reserved chained_ths) refs
+    val axioms =
+      maps (map (apsnd single)
+            o name_thm_pairs_from_ref ctxt reserved chained_ths) refs
   in
     case subgoal_count state of
       0 => priority "No subgoal!"
     | n =>
-      (kill_atps ();
-       priority (#2 (minimize_theorems params i n state name_thms_pairs)))
+      (kill_atps (); priority (#2 (minimize_theorems params i n state axioms)))
   end
 
 end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -67,11 +67,9 @@
    ("verbose", "false"),
    ("overlord", "false"),
    ("explicit_apply", "false"),
-   ("relevance_threshold", "40"),
-   ("relevance_convergence", "31"),
-   ("max_relevant_per_iter", "smart"),
+   ("relevance_thresholds", "45 85"),
+   ("max_relevant", "smart"),
    ("theory_relevant", "smart"),
-   ("defs_relevant", "false"),
    ("isar_proof", "false"),
    ("isar_shrink_factor", "1")]
 
@@ -84,7 +82,6 @@
    ("partial_types", "full_types"),
    ("implicit_apply", "explicit_apply"),
    ("theory_irrelevant", "theory_relevant"),
-   ("defs_irrelevant", "defs_relevant"),
    ("no_isar_proof", "isar_proof")]
 
 val params_for_minimize =
@@ -158,6 +155,14 @@
                     SOME n => n
                   | NONE => error ("Parameter " ^ quote name ^
                                    " must be assigned an integer value.")
+    fun lookup_int_pair name =
+      case lookup name of
+        NONE => (0, 0)
+      | SOME s => case s |> space_explode " " |> map Int.fromString of
+                    [SOME n1, SOME n2] => (n1, n2)
+                  | _ => error ("Parameter " ^ quote name ^
+                                "must be assigned a pair of integer values \
+                                \(e.g., \"60 95\")")
     fun lookup_int_option name =
       case lookup name of
         SOME "smart" => NONE
@@ -168,25 +173,20 @@
     val atps = lookup_string "atps" |> space_explode " "
     val full_types = lookup_bool "full_types"
     val explicit_apply = lookup_bool "explicit_apply"
-    val relevance_threshold =
-      0.01 * Real.fromInt (lookup_int "relevance_threshold")
-    val relevance_convergence =
-      0.01 * Real.fromInt (lookup_int "relevance_convergence")
-    val max_relevant_per_iter = lookup_int_option "max_relevant_per_iter"
+    val relevance_thresholds =
+      lookup_int_pair "relevance_thresholds"
+      |> pairself (fn n => 0.01 * Real.fromInt n)
+    val max_relevant = lookup_int_option "max_relevant"
     val theory_relevant = lookup_bool_option "theory_relevant"
-    val defs_relevant = lookup_bool "defs_relevant"
     val isar_proof = lookup_bool "isar_proof"
     val isar_shrink_factor = Int.max (1, lookup_int "isar_shrink_factor")
     val timeout = lookup_time "timeout"
   in
     {debug = debug, verbose = verbose, overlord = overlord, atps = atps,
      full_types = full_types, explicit_apply = explicit_apply,
-     relevance_threshold = relevance_threshold,
-     relevance_convergence = relevance_convergence,
-     max_relevant_per_iter = max_relevant_per_iter,
-     theory_relevant = theory_relevant, defs_relevant = defs_relevant,
-     isar_proof = isar_proof, isar_shrink_factor = isar_shrink_factor,
-     timeout = timeout}
+     relevance_thresholds = relevance_thresholds, max_relevant = max_relevant,
+     theory_relevant = theory_relevant, isar_proof = isar_proof,
+     isar_shrink_factor = isar_shrink_factor, timeout = timeout}
   end
 
 fun get_params thy = extract_params (default_raw_params thy)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -8,19 +8,18 @@
 
 signature SLEDGEHAMMER_PROOF_RECONSTRUCT =
 sig
+  type locality = Sledgehammer_Fact_Filter.locality
   type minimize_command = string list -> string
-
-  val metis_proof_text:
-    bool * minimize_command * string * (string * bool) vector * thm * int
-    -> string * (string * bool) list
-  val isar_proof_text:
+  type metis_params =
+    bool * minimize_command * string * (string * locality) list vector * thm
+    * int
+  type isar_params =
     string Symtab.table * bool * int * Proof.context * int list list
-    -> bool * minimize_command * string * (string * bool) vector * thm * int
-    -> string * (string * bool) list
-  val proof_text:
-    bool -> string Symtab.table * bool * int * Proof.context * int list list
-    -> bool * minimize_command * string * (string * bool) vector * thm * int
-    -> string * (string * bool) list
+  type text_result = string * (string * locality) list
+
+  val metis_proof_text : metis_params -> text_result
+  val isar_proof_text : isar_params -> metis_params -> text_result
+  val proof_text : bool -> isar_params -> metis_params -> text_result
 end;
 
 structure Sledgehammer_Proof_Reconstruct : SLEDGEHAMMER_PROOF_RECONSTRUCT =
@@ -33,6 +32,11 @@
 open Sledgehammer_Translate
 
 type minimize_command = string list -> string
+type metis_params =
+  bool * minimize_command * string * (string * locality) list vector * thm * int
+type isar_params =
+  string Symtab.table * bool * int * Proof.context * int list list
+type text_result = string * (string * locality) list
 
 (* Simple simplifications to ensure that sort annotations don't leave a trail of
    spurious "True"s. *)
@@ -60,7 +64,7 @@
 fun index_in_shape x = find_index (exists (curry (op =) x))
 fun is_axiom_number axiom_names num =
   num > 0 andalso num <= Vector.length axiom_names andalso
-  fst (Vector.sub (axiom_names, num - 1)) <> ""
+  not (null (Vector.sub (axiom_names, num - 1)))
 fun is_conjecture_number conjecture_shape num =
   index_in_shape num conjecture_shape >= 0
 
@@ -68,12 +72,12 @@
     Const (@{const_name Ex}, T) $ Abs (s, T', negate_term t')
   | negate_term (Const (@{const_name Ex}, T) $ Abs (s, T', t')) =
     Const (@{const_name All}, T) $ Abs (s, T', negate_term t')
-  | negate_term (@{const "op -->"} $ t1 $ t2) =
-    @{const "op &"} $ t1 $ negate_term t2
-  | negate_term (@{const "op &"} $ t1 $ t2) =
-    @{const "op |"} $ negate_term t1 $ negate_term t2
-  | negate_term (@{const "op |"} $ t1 $ t2) =
-    @{const "op &"} $ negate_term t1 $ negate_term t2
+  | negate_term (@{const HOL.implies} $ t1 $ t2) =
+    @{const HOL.conj} $ t1 $ negate_term t2
+  | negate_term (@{const HOL.conj} $ t1 $ t2) =
+    @{const HOL.disj} $ negate_term t1 $ negate_term t2
+  | negate_term (@{const HOL.disj} $ t1 $ t2) =
+    @{const HOL.conj} $ negate_term t1 $ negate_term t2
   | negate_term (@{const Not} $ t) = t
   | negate_term t = @{const Not} $ t
 
@@ -234,7 +238,7 @@
   fst o Scan.finite Symbol.stopper
             (Scan.error (!! (fn _ => raise Fail "unrecognized ATP output")
                             (parse_lines pool)))
-  o explode o strip_spaces
+  o explode o strip_spaces_except_between_ident_chars
 
 (**** INTERPRETATION OF TSTP SYNTAX TREES ****)
 
@@ -246,18 +250,18 @@
    constrained by information from type literals, or by type inference. *)
 fun type_from_fo_term tfrees (u as ATerm (a, us)) =
   let val Ts = map (type_from_fo_term tfrees) us in
-    case strip_prefix_and_undo_ascii type_const_prefix a of
+    case strip_prefix_and_unascii type_const_prefix a of
       SOME b => Type (invert_const b, Ts)
     | NONE =>
       if not (null us) then
         raise FO_TERM [u]  (* only "tconst"s have type arguments *)
-      else case strip_prefix_and_undo_ascii tfree_prefix a of
+      else case strip_prefix_and_unascii tfree_prefix a of
         SOME b =>
         let val s = "'" ^ b in
           TFree (s, AList.lookup (op =) tfrees s |> the_default HOLogic.typeS)
         end
       | NONE =>
-        case strip_prefix_and_undo_ascii tvar_prefix a of
+        case strip_prefix_and_unascii tvar_prefix a of
           SOME b => TVar (("'" ^ b, 0), HOLogic.typeS)
         | NONE =>
           (* Variable from the ATP, say "X1" *)
@@ -267,7 +271,7 @@
 (* Type class literal applied to a type. Returns triple of polarity, class,
    type. *)
 fun type_constraint_from_term pos tfrees (u as ATerm (a, us)) =
-  case (strip_prefix_and_undo_ascii class_prefix a,
+  case (strip_prefix_and_unascii class_prefix a,
         map (type_from_fo_term tfrees) us) of
     (SOME b, [T]) => (pos, b, T)
   | _ => raise FO_TERM [u]
@@ -309,9 +313,9 @@
             [typ_u, term_u] =>
             aux (SOME (type_from_fo_term tfrees typ_u)) extra_us term_u
           | _ => raise FO_TERM us
-        else case strip_prefix_and_undo_ascii const_prefix a of
+        else case strip_prefix_and_unascii const_prefix a of
           SOME "equal" =>
-          list_comb (Const (@{const_name "op ="}, HOLogic.typeT),
+          list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT),
                      map (aux NONE []) us)
         | SOME b =>
           let
@@ -341,10 +345,10 @@
             val ts = map (aux NONE []) (us @ extra_us)
             val T = map fastype_of ts ---> HOLogic.typeT
             val t =
-              case strip_prefix_and_undo_ascii fixed_var_prefix a of
+              case strip_prefix_and_unascii fixed_var_prefix a of
                 SOME b => Free (b, T)
               | NONE =>
-                case strip_prefix_and_undo_ascii schematic_var_prefix a of
+                case strip_prefix_and_unascii schematic_var_prefix a of
                   SOME b => Var ((b, 0), T)
                 | NONE =>
                   if is_tptp_variable a then
@@ -523,7 +527,7 @@
   | is_bad_free _ _ = false
 
 (* Vampire is keen on producing these. *)
-fun is_trivial_formula (@{const Not} $ (Const (@{const_name "op ="}, _)
+fun is_trivial_formula (@{const Not} $ (Const (@{const_name HOL.eq}, _)
                                         $ t1 $ t2)) = (t1 aconv t2)
   | is_trivial_formula _ = false
 
@@ -564,31 +568,31 @@
    "108. ... [input]". *)
 fun used_facts_in_atp_proof axiom_names atp_proof =
   let
-    fun axiom_name_at_index num =
+    fun axiom_names_at_index num =
       let val j = Int.fromString num |> the_default ~1 in
-        if is_axiom_number axiom_names j then
-          SOME (Vector.sub (axiom_names, j - 1))
-        else
-          NONE
+        if is_axiom_number axiom_names j then Vector.sub (axiom_names, j - 1)
+        else []
       end
     val tokens_of =
       String.tokens (fn c => not (Char.isAlphaNum c) andalso c <> #"_")
     fun do_line (tag :: num :: "axiom" :: (rest as _ :: _)) =
         if tag = "cnf" orelse tag = "fof" then
-          (case strip_prefix_and_undo_ascii axiom_prefix (List.last rest) of
+          (case strip_prefix_and_unascii axiom_prefix (List.last rest) of
              SOME name =>
              if member (op =) rest "file" then
-               SOME (name, is_true_for axiom_names name)
+               ([(name, name |> find_first_in_list_vector axiom_names |> the)]
+                handle Option.Option =>
+                       error ("No such fact: " ^ quote name ^ "."))
              else
-               axiom_name_at_index num
-           | NONE => axiom_name_at_index num)
+               axiom_names_at_index num
+           | NONE => axiom_names_at_index num)
         else
-          NONE
-      | do_line (num :: "0" :: "Inp" :: _) = axiom_name_at_index num
+          []
+      | do_line (num :: "0" :: "Inp" :: _) = axiom_names_at_index num
       | do_line (num :: rest) =
-        (case List.last rest of "input" => axiom_name_at_index num | _ => NONE)
-      | do_line _ = NONE
-  in atp_proof |> split_proof_lines |> map_filter (do_line o tokens_of) end
+        (case List.last rest of "input" => axiom_names_at_index num | _ => [])
+      | do_line _ = []
+  in atp_proof |> split_proof_lines |> maps (do_line o tokens_of) end
 
 val indent_size = 2
 val no_label = ("", ~1)
@@ -624,8 +628,8 @@
 
 fun used_facts axiom_names =
   used_facts_in_atp_proof axiom_names
-  #> List.partition snd
-  #> pairself (sort_distinct (string_ord) o map fst)
+  #> List.partition (curry (op =) Chained o snd)
+  #> pairself (sort_distinct (string_ord o pairself fst))
 
 fun metis_proof_text (full_types, minimize_command, atp_proof, axiom_names,
                       goal, i) =
@@ -633,9 +637,9 @@
     val (chained_lemmas, other_lemmas) = used_facts axiom_names atp_proof
     val n = Logic.count_prems (prop_of goal)
   in
-    (metis_line full_types i n other_lemmas ^
-     minimize_line minimize_command (other_lemmas @ chained_lemmas),
-     map (rpair false) other_lemmas @ map (rpair true) chained_lemmas)
+    (metis_line full_types i n (map fst other_lemmas) ^
+     minimize_line minimize_command (map fst (other_lemmas @ chained_lemmas)),
+     other_lemmas @ chained_lemmas)
   end
 
 (** Isar proof construction and manipulation **)
@@ -662,7 +666,7 @@
 
 fun add_fact_from_dep axiom_names num =
   if is_axiom_number axiom_names num then
-    apsnd (insert (op =) (fst (Vector.sub (axiom_names, num - 1))))
+    apsnd (union (op =) (map fst (Vector.sub (axiom_names, num - 1))))
   else
     apfst (insert (op =) (raw_prefix, num))
 
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -18,8 +18,8 @@
   val tfrees_name : string
   val prepare_problem :
     Proof.context -> bool -> bool -> bool -> bool -> term list -> term
-    -> ((string * bool) * thm) list
-    -> string problem * string Symtab.table * int * (string * bool) vector
+    -> ((string * 'a) * thm) list
+    -> string problem * string Symtab.table * int * (string * 'a) list vector
 end;
 
 structure Sledgehammer_Translate : SLEDGEHAMMER_TRANSLATE =
@@ -39,11 +39,11 @@
 (* Freshness almost guaranteed! *)
 val sledgehammer_weak_prefix = "Sledgehammer:"
 
-datatype fol_formula =
-  FOLFormula of {name: string,
-                 kind: kind,
-                 combformula: (name, combterm) formula,
-                 ctypes_sorts: typ list}
+type fol_formula =
+  {name: string,
+   kind: kind,
+   combformula: (name, combterm) formula,
+   ctypes_sorts: typ list}
 
 fun mk_anot phi = AConn (ANot, [phi])
 fun mk_aconn c phi1 phi2 = AConn (c, [phi1, phi2])
@@ -70,10 +70,10 @@
         do_quant bs AForall s T t'
       | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
         do_quant bs AExists s T t'
-      | @{const "op &"} $ t1 $ t2 => do_conn bs AAnd t1 t2
-      | @{const "op |"} $ t1 $ t2 => do_conn bs AOr t1 t2
-      | @{const "op -->"} $ t1 $ t2 => do_conn bs AImplies t1 t2
-      | Const (@{const_name "op ="}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
+      | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd t1 t2
+      | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr t1 t2
+      | @{const HOL.implies} $ t1 $ t2 => do_conn bs AImplies t1 t2
+      | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
         do_conn bs AIff t1 t2
       | _ => (fn ts => do_term bs (Envir.eta_contract t)
                        |>> AAtom ||> union (op =) ts)
@@ -97,7 +97,7 @@
       | aux j (t as Const (s, Type (_, [Type (_, [_, T']),
                                         Type (_, [_, res_T])]))
                     $ t2 $ Abs (var_s, var_T, t')) =
-        if s = @{const_name "op ="} orelse s = @{const_name "=="} then
+        if s = @{const_name HOL.eq} orelse s = @{const_name "=="} then
           let val var_t = Var ((var_s, j), var_T) in
             Const (s, T' --> T' --> res_T)
               $ betapply (t2, var_t) $ subst_bound (var_t, t')
@@ -125,10 +125,10 @@
             t0 $ Abs (s, T, aux (T :: Ts) t')
           | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
             aux Ts (t0 $ eta_expand Ts t1 1)
-          | (t0 as @{const "op &"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
-          | (t0 as @{const "op |"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
-          | (t0 as @{const "op -->"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
-          | (t0 as Const (@{const_name "op ="}, Type (_, [@{typ bool}, _])))
+          | (t0 as @{const HOL.conj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
+          | (t0 as @{const HOL.disj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
+          | (t0 as @{const HOL.implies}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
+          | (t0 as Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])))
               $ t1 $ t2 =>
             t0 $ aux Ts t1 $ aux Ts t2
           | _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then
@@ -190,15 +190,14 @@
               |> kind <> Axiom ? freeze_term
     val (combformula, ctypes_sorts) = combformula_for_prop thy t []
   in
-    FOLFormula {name = name, combformula = combformula, kind = kind,
-                ctypes_sorts = ctypes_sorts}
+    {name = name, combformula = combformula, kind = kind,
+     ctypes_sorts = ctypes_sorts}
   end
 
-fun make_axiom ctxt presimp ((name, chained), th) =
+fun make_axiom ctxt presimp ((name, loc), th) =
   case make_formula ctxt presimp name Axiom (prop_of th) of
-    FOLFormula {combformula = AAtom (CombConst (("c_True", _), _, _)), ...} =>
-    NONE
-  | formula => SOME ((name, chained), formula)
+    {combformula = AAtom (CombConst (("c_True", _), _, _)), ...} => NONE
+  | formula => SOME ((name, loc), formula)
 fun make_conjecture ctxt ts =
   let val last = length ts - 1 in
     map2 (fn j => make_formula ctxt true (Int.toString j)
@@ -215,12 +214,14 @@
 fun count_combformula (AQuant (_, _, phi)) = count_combformula phi
   | count_combformula (AConn (_, phis)) = fold count_combformula phis
   | count_combformula (AAtom tm) = count_combterm tm
-fun count_fol_formula (FOLFormula {combformula, ...}) =
+fun count_fol_formula ({combformula, ...} : fol_formula) =
   count_combformula combformula
 
 val optional_helpers =
-  [(["c_COMBI", "c_COMBK"], @{thms COMBI_def COMBK_def}),
-   (["c_COMBB", "c_COMBC"], @{thms COMBB_def COMBC_def}),
+  [(["c_COMBI"], @{thms COMBI_def}),
+   (["c_COMBK"], @{thms COMBK_def}),
+   (["c_COMBB"], @{thms COMBB_def}),
+   (["c_COMBC"], @{thms COMBC_def}),
    (["c_COMBS"], @{thms COMBS_def})]
 val optional_typed_helpers =
   [(["c_True", "c_False", "c_If"], @{thms True_or_False}),
@@ -272,7 +273,7 @@
     val (supers', arity_clauses) = make_arity_clauses thy tycons supers
     val class_rel_clauses = make_class_rel_clauses thy subs supers'
   in
-    (Vector.fromList axiom_names,
+    (axiom_names |> map single |> Vector.fromList,
      (conjectures, axioms, helper_facts, class_rel_clauses, arity_clauses))
   end
 
@@ -326,13 +327,13 @@
       | aux (AAtom tm) = AAtom (fo_term_for_combterm full_types tm)
   in aux end
 
-fun formula_for_axiom full_types (FOLFormula {combformula, ctypes_sorts, ...}) =
+fun formula_for_axiom full_types
+                      ({combformula, ctypes_sorts, ...} : fol_formula) =
   mk_ahorn (map (formula_for_fo_literal o fo_literal_for_type_literal)
                 (type_literals_for_types ctypes_sorts))
            (formula_for_combformula full_types combformula)
 
-fun problem_line_for_fact prefix full_types
-                          (formula as FOLFormula {name, kind, ...}) =
+fun problem_line_for_fact prefix full_types (formula as {name, kind, ...}) =
   Fof (prefix ^ ascii_of name, kind, formula_for_axiom full_types formula)
 
 fun problem_line_for_class_rel_clause (ClassRelClause {name, subclass,
@@ -357,11 +358,11 @@
                      (fo_literal_for_arity_literal conclLit)))
 
 fun problem_line_for_conjecture full_types
-                                (FOLFormula {name, kind, combformula, ...}) =
+                                ({name, kind, combformula, ...} : fol_formula) =
   Fof (conjecture_prefix ^ name, kind,
        formula_for_combformula full_types combformula)
 
-fun free_type_literals_for_conjecture (FOLFormula {ctypes_sorts, ...}) =
+fun free_type_literals_for_conjecture ({ctypes_sorts, ...} : fol_formula) =
   map fo_literal_for_type_literal (type_literals_for_types ctypes_sorts)
 
 fun problem_line_for_free_type lit =
@@ -407,7 +408,7 @@
        16383 (* large number *)
      else if full_types then
        0
-     else case strip_prefix_and_undo_ascii const_prefix s of
+     else case strip_prefix_and_unascii const_prefix s of
        SOME s' => num_type_args thy (invert_const s')
      | NONE => 0)
   | min_arity_of _ _ (SOME the_const_tab) s =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -6,10 +6,11 @@
 
 signature SLEDGEHAMMER_UTIL =
 sig
-  val is_true_for : (string * bool) vector -> string -> bool
+  val find_first_in_list_vector : (''a * 'b) list vector -> ''a -> 'b option
   val plural_s : int -> string
   val serial_commas : string -> string list -> string list
-  val strip_spaces : string -> string
+  val simplify_spaces : string -> string
+  val strip_spaces_except_between_ident_chars : string -> string
   val parse_bool_option : bool -> string -> string -> bool option
   val parse_time_option : string -> string -> Time.time option
   val scan_integer : string list -> int * string list
@@ -28,8 +29,9 @@
 structure Sledgehammer_Util : SLEDGEHAMMER_UTIL =
 struct
 
-fun is_true_for v s =
-  Vector.foldl (fn ((s', b'), b) => if s' = s then b' else b) false v
+fun find_first_in_list_vector vec key =
+  Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key
+                 | (_, value) => value) NONE vec
 
 fun plural_s n = if n = 1 then "" else "s"
 
@@ -39,24 +41,27 @@
   | serial_commas conj [s1, s2, s3] = [s1 ^ ",", s2 ^ ",", conj, s3]
   | serial_commas conj (s :: ss) = s ^ "," :: serial_commas conj ss
 
-fun is_ident_char c = Char.isAlphaNum c orelse c = #"_"
-
-fun strip_spaces_in_list [] = ""
-  | strip_spaces_in_list [c1] = if Char.isSpace c1 then "" else str c1
-  | strip_spaces_in_list [c1, c2] =
-    strip_spaces_in_list [c1] ^ strip_spaces_in_list [c2]
-  | strip_spaces_in_list (c1 :: c2 :: c3 :: cs) =
+fun strip_spaces_in_list _ [] = ""
+  | strip_spaces_in_list _ [c1] = if Char.isSpace c1 then "" else str c1
+  | strip_spaces_in_list is_evil [c1, c2] =
+    strip_spaces_in_list is_evil [c1] ^ strip_spaces_in_list is_evil [c2]
+  | strip_spaces_in_list is_evil (c1 :: c2 :: c3 :: cs) =
     if Char.isSpace c1 then
-      strip_spaces_in_list (c2 :: c3 :: cs)
+      strip_spaces_in_list is_evil (c2 :: c3 :: cs)
     else if Char.isSpace c2 then
       if Char.isSpace c3 then
-        strip_spaces_in_list (c1 :: c3 :: cs)
+        strip_spaces_in_list is_evil (c1 :: c3 :: cs)
       else
-        str c1 ^ (if forall is_ident_char [c1, c3] then " " else "") ^
-        strip_spaces_in_list (c3 :: cs)
+        str c1 ^ (if forall is_evil [c1, c3] then " " else "") ^
+        strip_spaces_in_list is_evil (c3 :: cs)
     else
-      str c1 ^ strip_spaces_in_list (c2 :: c3 :: cs)
-val strip_spaces = strip_spaces_in_list o String.explode
+      str c1 ^ strip_spaces_in_list is_evil (c2 :: c3 :: cs)
+fun strip_spaces is_evil = strip_spaces_in_list is_evil o String.explode
+
+val simplify_spaces = strip_spaces (K true)
+
+fun is_ident_char c = Char.isAlphaNum c orelse c = #"_"
+val strip_spaces_except_between_ident_chars = strip_spaces is_ident_char
 
 fun parse_bool_option option name s =
   (case s of
--- a/src/HOL/Tools/TFL/dcterm.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/TFL/dcterm.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -79,11 +79,11 @@
   in capply c (uncurry cabs r) end;
 
 
-local val c = mk_hol_const(@{const_name "op &"}, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
+local val c = mk_hol_const(@{const_name HOL.conj}, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
 in fun mk_conj(conj1,conj2) = capply (capply c conj1) conj2
 end;
 
-local val c = mk_hol_const(@{const_name "op |"}, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
+local val c = mk_hol_const(@{const_name HOL.disj}, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
 in fun mk_disj(disj1,disj2) = capply (capply c disj1) disj2
 end;
 
@@ -127,10 +127,10 @@
 
 val dest_neg    = dest_monop @{const_name Not}
 val dest_pair   = dest_binop @{const_name Pair}
-val dest_eq     = dest_binop @{const_name "op ="}
-val dest_imp    = dest_binop @{const_name "op -->"}
-val dest_conj   = dest_binop @{const_name "op &"}
-val dest_disj   = dest_binop @{const_name "op |"}
+val dest_eq     = dest_binop @{const_name HOL.eq}
+val dest_imp    = dest_binop @{const_name HOL.implies}
+val dest_conj   = dest_binop @{const_name HOL.conj}
+val dest_disj   = dest_binop @{const_name HOL.disj}
 val dest_select = dest_binder @{const_name Eps}
 val dest_exists = dest_binder @{const_name Ex}
 val dest_forall = dest_binder @{const_name All}
--- a/src/HOL/Tools/TFL/post.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/TFL/post.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -67,7 +67,7 @@
 val P_imp_P_eq_True = @{thm eqTrueI} RS eq_reflection;
 fun mk_meta_eq r = case concl_of r of
      Const("==",_)$_$_ => r
-  |   _ $(Const(@{const_name "op ="},_)$_$_) => r RS eq_reflection
+  |   _ $(Const(@{const_name HOL.eq},_)$_$_) => r RS eq_reflection
   |   _ => r RS P_imp_P_eq_True
 
 (*Is this the best way to invoke the simplifier??*)
--- a/src/HOL/Tools/TFL/rules.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/TFL/rules.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -46,10 +46,6 @@
   val simpl_conv: simpset -> thm list -> cterm -> thm
 
   val rbeta: thm -> thm
-(* For debugging my isabelle solver in the conditional rewriter *)
-  val term_ref: term list Unsynchronized.ref
-  val thm_ref: thm list Unsynchronized.ref
-  val ss_ref: simpset list Unsynchronized.ref
   val tracing: bool Unsynchronized.ref
   val CONTEXT_REWRITE_RULE: term * term list * thm * thm list
                              -> thm -> thm * term list
@@ -542,9 +538,6 @@
 (*---------------------------------------------------------------------------
  * Trace information for the rewriter
  *---------------------------------------------------------------------------*)
-val term_ref = Unsynchronized.ref [] : term list Unsynchronized.ref
-val ss_ref = Unsynchronized.ref [] : simpset list Unsynchronized.ref;
-val thm_ref = Unsynchronized.ref [] : thm list Unsynchronized.ref;
 val tracing = Unsynchronized.ref false;
 
 fun say s = if !tracing then writeln s else ();
@@ -665,9 +658,6 @@
      val ss0 = Simplifier.global_context (Thm.theory_of_thm th) empty_ss
      val pbeta_reduce = simpl_conv ss0 [@{thm split_conv} RS eq_reflection];
      val tc_list = Unsynchronized.ref []: term list Unsynchronized.ref
-     val dummy = term_ref := []
-     val dummy = thm_ref  := []
-     val dummy = ss_ref  := []
      val cut_lemma' = cut_lemma RS eq_reflection
      fun prover used ss thm =
      let fun cong_prover ss thm =
@@ -676,8 +666,6 @@
              val dummy = print_thms "cntxt:" cntxt
              val dummy = say "cong rule:"
              val dummy = say (Display.string_of_thm_without_context thm)
-             val dummy = thm_ref := (thm :: !thm_ref)
-             val dummy = ss_ref := (ss :: !ss_ref)
              (* Unquantified eliminate *)
              fun uq_eliminate (thm,imp,thy) =
                  let val tych = cterm_of thy
@@ -784,7 +772,6 @@
               val dummy = tc_list := (TC :: !tc_list)
               val nestedp = is_some (S.find_term is_func TC)
               val dummy = if nestedp then say "nested" else say "not_nested"
-              val dummy = term_ref := ([func,TC]@(!term_ref))
               val th' = if nestedp then raise RULES_ERR "solver" "nested function"
                         else let val cTC = cterm_of thy
                                               (HOLogic.mk_Trueprop TC)
--- a/src/HOL/Tools/TFL/usyntax.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/TFL/usyntax.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -159,7 +159,7 @@
 
 
 fun mk_imp{ant,conseq} =
-   let val c = Const(@{const_name "op -->"},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
+   let val c = Const(@{const_name HOL.implies},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
    in list_comb(c,[ant,conseq])
    end;
 
@@ -183,12 +183,12 @@
 
 
 fun mk_conj{conj1,conj2} =
-   let val c = Const(@{const_name "op &"},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
+   let val c = Const(@{const_name HOL.conj},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
    in list_comb(c,[conj1,conj2])
    end;
 
 fun mk_disj{disj1,disj2} =
-   let val c = Const(@{const_name "op |"},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
+   let val c = Const(@{const_name HOL.disj},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
    in list_comb(c,[disj1,disj2])
    end;
 
@@ -244,10 +244,10 @@
      end
   | dest_abs _ _ =  raise USYN_ERR "dest_abs" "not an abstraction";
 
-fun dest_eq(Const(@{const_name "op ="},_) $ M $ N) = {lhs=M, rhs=N}
+fun dest_eq(Const(@{const_name HOL.eq},_) $ M $ N) = {lhs=M, rhs=N}
   | dest_eq _ = raise USYN_ERR "dest_eq" "not an equality";
 
-fun dest_imp(Const(@{const_name "op -->"},_) $ M $ N) = {ant=M, conseq=N}
+fun dest_imp(Const(@{const_name HOL.implies},_) $ M $ N) = {ant=M, conseq=N}
   | dest_imp _ = raise USYN_ERR "dest_imp" "not an implication";
 
 fun dest_forall(Const(@{const_name All},_) $ (a as Abs _)) = fst (dest_abs [] a)
@@ -259,10 +259,10 @@
 fun dest_neg(Const(@{const_name Not},_) $ M) = M
   | dest_neg _ = raise USYN_ERR "dest_neg" "not a negation";
 
-fun dest_conj(Const(@{const_name "op &"},_) $ M $ N) = {conj1=M, conj2=N}
+fun dest_conj(Const(@{const_name HOL.conj},_) $ M $ N) = {conj1=M, conj2=N}
   | dest_conj _ = raise USYN_ERR "dest_conj" "not a conjunction";
 
-fun dest_disj(Const(@{const_name "op |"},_) $ M $ N) = {disj1=M, disj2=N}
+fun dest_disj(Const(@{const_name HOL.disj},_) $ M $ N) = {disj1=M, disj2=N}
   | dest_disj _ = raise USYN_ERR "dest_disj" "not a disjunction";
 
 fun mk_pair{fst,snd} =
--- a/src/HOL/Tools/choice_specification.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/choice_specification.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -220,8 +220,8 @@
                     |> process_all (zip3 alt_names rew_imps frees)
             end
 
-      fun after_qed [[thm]] = (ProofContext.theory (fn thy =>
-        #1 (post_process (add_specification axiomatic (zip3 names cnames overloaded) (thy, thm)))));
+      fun after_qed [[thm]] = ProofContext.background_theory (fn thy =>
+        #1 (post_process (add_specification axiomatic (zip3 names cnames overloaded) (thy, thm))));
     in
       thy
       |> ProofContext.init_global
--- a/src/HOL/Tools/cnf_funcs.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/cnf_funcs.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -95,17 +95,17 @@
 
 fun is_atom (Const (@{const_name False}, _))                                           = false
   | is_atom (Const (@{const_name True}, _))                                            = false
-  | is_atom (Const (@{const_name "op &"}, _) $ _ $ _)                                    = false
-  | is_atom (Const (@{const_name "op |"}, _) $ _ $ _)                                    = false
-  | is_atom (Const (@{const_name "op -->"}, _) $ _ $ _)                                  = false
-  | is_atom (Const (@{const_name "op ="}, Type ("fun", @{typ bool} :: _)) $ _ $ _)       = false
+  | is_atom (Const (@{const_name HOL.conj}, _) $ _ $ _)                                    = false
+  | is_atom (Const (@{const_name HOL.disj}, _) $ _ $ _)                                    = false
+  | is_atom (Const (@{const_name HOL.implies}, _) $ _ $ _)                                  = false
+  | is_atom (Const (@{const_name HOL.eq}, Type ("fun", @{typ bool} :: _)) $ _ $ _)       = false
   | is_atom (Const (@{const_name Not}, _) $ _)                                         = false
   | is_atom _                                                              = true;
 
 fun is_literal (Const (@{const_name Not}, _) $ x) = is_atom x
   | is_literal x                      = is_atom x;
 
-fun is_clause (Const (@{const_name "op |"}, _) $ x $ y) = is_clause x andalso is_clause y
+fun is_clause (Const (@{const_name HOL.disj}, _) $ x $ y) = is_clause x andalso is_clause y
   | is_clause x                           = is_literal x;
 
 (* ------------------------------------------------------------------------- *)
@@ -184,28 +184,28 @@
 
 (* Theory.theory -> Term.term -> Thm.thm *)
 
-fun make_nnf_thm thy (Const (@{const_name "op &"}, _) $ x $ y) =
+fun make_nnf_thm thy (Const (@{const_name HOL.conj}, _) $ x $ y) =
 	let
 		val thm1 = make_nnf_thm thy x
 		val thm2 = make_nnf_thm thy y
 	in
 		conj_cong OF [thm1, thm2]
 	end
-  | make_nnf_thm thy (Const (@{const_name "op |"}, _) $ x $ y) =
+  | make_nnf_thm thy (Const (@{const_name HOL.disj}, _) $ x $ y) =
 	let
 		val thm1 = make_nnf_thm thy x
 		val thm2 = make_nnf_thm thy y
 	in
 		disj_cong OF [thm1, thm2]
 	end
-  | make_nnf_thm thy (Const (@{const_name "op -->"}, _) $ x $ y) =
+  | make_nnf_thm thy (Const (@{const_name HOL.implies}, _) $ x $ y) =
 	let
 		val thm1 = make_nnf_thm thy (HOLogic.Not $ x)
 		val thm2 = make_nnf_thm thy y
 	in
 		make_nnf_imp OF [thm1, thm2]
 	end
-  | make_nnf_thm thy (Const (@{const_name "op ="}, Type ("fun", @{typ bool} :: _)) $ x $ y) =
+  | make_nnf_thm thy (Const (@{const_name HOL.eq}, Type ("fun", @{typ bool} :: _)) $ x $ y) =
 	let
 		val thm1 = make_nnf_thm thy x
 		val thm2 = make_nnf_thm thy (HOLogic.Not $ x)
@@ -218,28 +218,28 @@
 	make_nnf_not_false
   | make_nnf_thm thy (Const (@{const_name Not}, _) $ Const (@{const_name True}, _)) =
 	make_nnf_not_true
-  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name "op &"}, _) $ x $ y)) =
+  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.conj}, _) $ x $ y)) =
 	let
 		val thm1 = make_nnf_thm thy (HOLogic.Not $ x)
 		val thm2 = make_nnf_thm thy (HOLogic.Not $ y)
 	in
 		make_nnf_not_conj OF [thm1, thm2]
 	end
-  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name "op |"}, _) $ x $ y)) =
+  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.disj}, _) $ x $ y)) =
 	let
 		val thm1 = make_nnf_thm thy (HOLogic.Not $ x)
 		val thm2 = make_nnf_thm thy (HOLogic.Not $ y)
 	in
 		make_nnf_not_disj OF [thm1, thm2]
 	end
-  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name "op -->"}, _) $ x $ y)) =
+  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.implies}, _) $ x $ y)) =
 	let
 		val thm1 = make_nnf_thm thy x
 		val thm2 = make_nnf_thm thy (HOLogic.Not $ y)
 	in
 		make_nnf_not_imp OF [thm1, thm2]
 	end
-  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name "op ="}, Type ("fun", @{typ bool} :: _)) $ x $ y)) =
+  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.eq}, Type ("fun", @{typ bool} :: _)) $ x $ y)) =
 	let
 		val thm1 = make_nnf_thm thy x
 		val thm2 = make_nnf_thm thy (HOLogic.Not $ x)
@@ -276,7 +276,7 @@
 
 (* Theory.theory -> Term.term -> Thm.thm *)
 
-fun simp_True_False_thm thy (Const (@{const_name "op &"}, _) $ x $ y) =
+fun simp_True_False_thm thy (Const (@{const_name HOL.conj}, _) $ x $ y) =
 	let
 		val thm1 = simp_True_False_thm thy x
 		val x'   = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm1
@@ -298,7 +298,7 @@
 					conj_cong OF [thm1, thm2]  (* (x & y) = (x' & y') *)
 			end
 	end
-  | simp_True_False_thm thy (Const (@{const_name "op |"}, _) $ x $ y) =
+  | simp_True_False_thm thy (Const (@{const_name HOL.disj}, _) $ x $ y) =
 	let
 		val thm1 = simp_True_False_thm thy x
 		val x'   = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm1
@@ -334,24 +334,24 @@
 fun make_cnf_thm thy t =
 let
 	(* Term.term -> Thm.thm *)
-	fun make_cnf_thm_from_nnf (Const (@{const_name "op &"}, _) $ x $ y) =
+	fun make_cnf_thm_from_nnf (Const (@{const_name HOL.conj}, _) $ x $ y) =
 		let
 			val thm1 = make_cnf_thm_from_nnf x
 			val thm2 = make_cnf_thm_from_nnf y
 		in
 			conj_cong OF [thm1, thm2]
 		end
-	  | make_cnf_thm_from_nnf (Const (@{const_name "op |"}, _) $ x $ y) =
+	  | make_cnf_thm_from_nnf (Const (@{const_name HOL.disj}, _) $ x $ y) =
 		let
 			(* produces a theorem "(x' | y') = t'", where x', y', and t' are in CNF *)
-			fun make_cnf_disj_thm (Const (@{const_name "op &"}, _) $ x1 $ x2) y' =
+			fun make_cnf_disj_thm (Const (@{const_name HOL.conj}, _) $ x1 $ x2) y' =
 				let
 					val thm1 = make_cnf_disj_thm x1 y'
 					val thm2 = make_cnf_disj_thm x2 y'
 				in
 					make_cnf_disj_conj_l OF [thm1, thm2]  (* ((x1 & x2) | y') = ((x1 | y')' & (x2 | y')') *)
 				end
-			  | make_cnf_disj_thm x' (Const (@{const_name "op &"}, _) $ y1 $ y2) =
+			  | make_cnf_disj_thm x' (Const (@{const_name HOL.conj}, _) $ y1 $ y2) =
 				let
 					val thm1 = make_cnf_disj_thm x' y1
 					val thm2 = make_cnf_disj_thm x' y2
@@ -403,27 +403,27 @@
 	val var_id = Unsynchronized.ref 0  (* properly initialized below *)
 	fun new_free () =
 		Free ("cnfx_" ^ string_of_int (Unsynchronized.inc var_id), HOLogic.boolT)
-	fun make_cnfx_thm_from_nnf (Const (@{const_name "op &"}, _) $ x $ y) : thm =
+	fun make_cnfx_thm_from_nnf (Const (@{const_name HOL.conj}, _) $ x $ y) : thm =
 		let
 			val thm1 = make_cnfx_thm_from_nnf x
 			val thm2 = make_cnfx_thm_from_nnf y
 		in
 			conj_cong OF [thm1, thm2]
 		end
-	  | make_cnfx_thm_from_nnf (Const (@{const_name "op |"}, _) $ x $ y) =
+	  | make_cnfx_thm_from_nnf (Const (@{const_name HOL.disj}, _) $ x $ y) =
 		if is_clause x andalso is_clause y then
 			inst_thm thy [HOLogic.mk_disj (x, y)] iff_refl
 		else if is_literal y orelse is_literal x then let
 			(* produces a theorem "(x' | y') = t'", where x', y', and t' are *)
 			(* almost in CNF, and x' or y' is a literal                      *)
-			fun make_cnfx_disj_thm (Const (@{const_name "op &"}, _) $ x1 $ x2) y' =
+			fun make_cnfx_disj_thm (Const (@{const_name HOL.conj}, _) $ x1 $ x2) y' =
 				let
 					val thm1 = make_cnfx_disj_thm x1 y'
 					val thm2 = make_cnfx_disj_thm x2 y'
 				in
 					make_cnf_disj_conj_l OF [thm1, thm2]  (* ((x1 & x2) | y') = ((x1 | y')' & (x2 | y')') *)
 				end
-			  | make_cnfx_disj_thm x' (Const (@{const_name "op &"}, _) $ y1 $ y2) =
+			  | make_cnfx_disj_thm x' (Const (@{const_name HOL.conj}, _) $ y1 $ y2) =
 				let
 					val thm1 = make_cnfx_disj_thm x' y1
 					val thm2 = make_cnfx_disj_thm x' y2
--- a/src/HOL/Tools/groebner.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/groebner.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -395,7 +395,7 @@
 
 fun refute_disj rfn tm =
  case term_of tm of
-  Const(@{const_name "op |"},_)$l$r =>
+  Const(@{const_name HOL.disj},_)$l$r =>
    compose_single(refute_disj rfn (dest_arg tm),2,compose_single(refute_disj rfn (dest_arg1 tm),2,disjE))
   | _ => rfn tm ;
 
@@ -409,7 +409,7 @@
     | _  => false;
 fun is_eq t =
  case term_of t of
- (Const(@{const_name "op ="},_)$_$_) => true
+ (Const(@{const_name HOL.eq},_)$_$_) => true
 | _  => false;
 
 fun end_itlist f l =
@@ -454,7 +454,7 @@
 val specl = fold_rev (fn x => fn th => instantiate' [] [SOME x] (th RS spec));
 
 val cTrp = @{cterm "Trueprop"};
-val cConj = @{cterm "op &"};
+val cConj = @{cterm HOL.conj};
 val (cNot,false_tm) = (@{cterm "Not"}, @{cterm "False"});
 val assume_Trueprop = mk_comb cTrp #> assume;
 val list_mk_conj = list_mk_binop cConj;
@@ -479,22 +479,22 @@
   (* FIXME : copied from cqe.ML -- complex QE*)
 fun conjuncts ct =
  case term_of ct of
-  @{term "op &"}$_$_ => (Thm.dest_arg1 ct)::(conjuncts (Thm.dest_arg ct))
+  @{term HOL.conj}$_$_ => (Thm.dest_arg1 ct)::(conjuncts (Thm.dest_arg ct))
 | _ => [ct];
 
 fun fold1 f = foldr1 (uncurry f);
 
-val list_conj = fold1 (fn c => fn c' => Thm.capply (Thm.capply @{cterm "op &"} c) c') ;
+val list_conj = fold1 (fn c => fn c' => Thm.capply (Thm.capply @{cterm HOL.conj} c) c') ;
 
 fun mk_conj_tab th = 
  let fun h acc th = 
    case prop_of th of
-   @{term "Trueprop"}$(@{term "op &"}$p$q) => 
+   @{term "Trueprop"}$(@{term HOL.conj}$p$q) => 
      h (h acc (th RS conjunct2)) (th RS conjunct1)
   | @{term "Trueprop"}$p => (p,th)::acc
 in fold (Termtab.insert Thm.eq_thm) (h [] th) Termtab.empty end;
 
-fun is_conj (@{term "op &"}$_$_) = true
+fun is_conj (@{term HOL.conj}$_$_) = true
   | is_conj _ = false;
 
 fun prove_conj tab cjs = 
@@ -852,14 +852,14 @@
  fun unwind_polys_conv tm =
  let 
   val (vars,bod) = strip_exists tm
-  val cjs = striplist (dest_binary @{cterm "op &"}) bod
+  val cjs = striplist (dest_binary @{cterm HOL.conj}) bod
   val th1 = (the (get_first (try (isolate_variable vars)) cjs) 
              handle Option => raise CTERM ("unwind_polys_conv",[tm]))
   val eq = Thm.lhs_of th1
-  val bod' = list_mk_binop @{cterm "op &"} (eq::(remove op aconvc eq cjs))
+  val bod' = list_mk_binop @{cterm HOL.conj} (eq::(remove op aconvc eq cjs))
   val th2 = conj_ac_rule (mk_eq bod bod')
   val th3 = transitive th2 
-         (Drule.binop_cong_rule @{cterm "op &"} th1 
+         (Drule.binop_cong_rule @{cterm HOL.conj} th1 
                 (reflexive (Thm.dest_arg (Thm.rhs_of th2))))
   val v = Thm.dest_arg1(Thm.dest_arg1(Thm.rhs_of th3))
   val vars' = (remove op aconvc v vars) @ [v]
@@ -923,15 +923,15 @@
 
 fun find_term bounds tm =
   (case term_of tm of
-    Const (@{const_name "op ="}, T) $ _ $ _ =>
+    Const (@{const_name HOL.eq}, T) $ _ $ _ =>
       if domain_type T = HOLogic.boolT then find_args bounds tm
       else dest_arg tm
   | Const (@{const_name Not}, _) $ _ => find_term bounds (dest_arg tm)
   | Const (@{const_name All}, _) $ _ => find_body bounds (dest_arg tm)
   | Const (@{const_name Ex}, _) $ _ => find_body bounds (dest_arg tm)
-  | Const (@{const_name "op &"}, _) $ _ $ _ => find_args bounds tm
-  | Const (@{const_name "op |"}, _) $ _ $ _ => find_args bounds tm
-  | Const (@{const_name "op -->"}, _) $ _ $ _ => find_args bounds tm
+  | Const (@{const_name HOL.conj}, _) $ _ $ _ => find_args bounds tm
+  | Const (@{const_name HOL.disj}, _) $ _ $ _ => find_args bounds tm
+  | Const (@{const_name HOL.implies}, _) $ _ $ _ => find_args bounds tm
   | @{term "op ==>"} $_$_ => find_args bounds tm
   | Const("op ==",_)$_$_ => find_args bounds tm
   | @{term Trueprop}$_ => find_term bounds (dest_arg tm)
@@ -985,7 +985,7 @@
 
 local
  fun lhs t = case term_of t of
-  Const(@{const_name "op ="},_)$_$_ => Thm.dest_arg1 t
+  Const(@{const_name HOL.eq},_)$_$_ => Thm.dest_arg1 t
  | _=> raise CTERM ("ideal_tac - lhs",[t])
  fun exitac NONE = no_tac
    | exitac (SOME y) = rtac (instantiate' [SOME (ctyp_of_term y)] [NONE,SOME y] exI) 1
--- a/src/HOL/Tools/hologic.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/hologic.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -49,7 +49,7 @@
   val exists_const: typ -> term
   val mk_exists: string * typ * term -> term
   val choice_const: typ -> term
-  val class_eq: string
+  val class_equal: string
   val mk_binop: string -> term * term -> term
   val mk_binrel: string -> term * term -> term
   val dest_bin: string -> typ -> term -> term * term
@@ -208,38 +208,38 @@
   let val (th1, th2) = conj_elim th
   in conj_elims th1 @ conj_elims th2 end handle THM _ => [th];
 
-val conj = @{term "op &"}
-and disj = @{term "op |"}
-and imp = @{term "op -->"}
-and Not = @{term "Not"};
+val conj = @{term HOL.conj}
+and disj = @{term HOL.disj}
+and imp = @{term implies}
+and Not = @{term Not};
 
 fun mk_conj (t1, t2) = conj $ t1 $ t2
 and mk_disj (t1, t2) = disj $ t1 $ t2
 and mk_imp (t1, t2) = imp $ t1 $ t2
 and mk_not t = Not $ t;
 
-fun dest_conj (Const ("op &", _) $ t $ t') = t :: dest_conj t'
+fun dest_conj (Const ("HOL.conj", _) $ t $ t') = t :: dest_conj t'
   | dest_conj t = [t];
 
-fun dest_disj (Const ("op |", _) $ t $ t') = t :: dest_disj t'
+fun dest_disj (Const ("HOL.disj", _) $ t $ t') = t :: dest_disj t'
   | dest_disj t = [t];
 
 (*Like dest_disj, but flattens disjunctions however nested*)
-fun disjuncts_aux (Const ("op |", _) $ t $ t') disjs = disjuncts_aux t (disjuncts_aux t' disjs)
+fun disjuncts_aux (Const ("HOL.disj", _) $ t $ t') disjs = disjuncts_aux t (disjuncts_aux t' disjs)
   | disjuncts_aux t disjs = t::disjs;
 
 fun disjuncts t = disjuncts_aux t [];
 
-fun dest_imp (Const ("op -->", _) $ A $ B) = (A, B)
+fun dest_imp (Const ("HOL.implies", _) $ A $ B) = (A, B)
   | dest_imp  t = raise TERM ("dest_imp", [t]);
 
 fun dest_not (Const ("HOL.Not", _) $ t) = t
   | dest_not t = raise TERM ("dest_not", [t]);
 
-fun eq_const T = Const ("op =", T --> T --> boolT);
+fun eq_const T = Const ("HOL.eq", T --> T --> boolT);
 fun mk_eq (t, u) = eq_const (fastype_of t) $ t $ u;
 
-fun dest_eq (Const ("op =", _) $ lhs $ rhs) = (lhs, rhs)
+fun dest_eq (Const ("HOL.eq", _) $ lhs $ rhs) = (lhs, rhs)
   | dest_eq t = raise TERM ("dest_eq", [t])
 
 fun all_const T = Const ("HOL.All", [T --> boolT] ---> boolT);
@@ -251,7 +251,7 @@
 
 fun choice_const T = Const("Hilbert_Choice.Eps", (T --> boolT) --> T);
 
-val class_eq = "HOL.eq";
+val class_equal = "HOL.equal";
 
 
 (* binary operations and relations *)
--- a/src/HOL/Tools/inductive.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/inductive.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -182,7 +182,7 @@
   in
     case concl_of thm of
       Const ("==", _) $ _ $ _ => eq2mono (thm RS meta_eq_to_obj_eq)
-    | _ $ (Const (@{const_name "op ="}, _) $ _ $ _) => eq2mono thm
+    | _ $ (Const (@{const_name HOL.eq}, _) $ _ $ _) => eq2mono thm
     | _ $ (Const (@{const_name Orderings.less_eq}, _) $ _ $ _) =>
       dest_less_concl (Seq.hd (REPEAT (FIRSTGOAL
         (resolve_tac [@{thm le_funI}, @{thm le_boolI'}])) thm))
--- a/src/HOL/Tools/inductive_codegen.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/inductive_codegen.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -52,7 +52,7 @@
     fun thyname_of s = (case optmod of
       NONE => Codegen.thyname_of_const thy s | SOME s => s);
   in (case Option.map strip_comb (try HOLogic.dest_Trueprop (concl_of thm)) of
-      SOME (Const (@{const_name "op ="}, _), [t, _]) =>
+      SOME (Const (@{const_name HOL.eq}, _), [t, _]) =>
         (case head_of t of
           Const (s, _) =>
             CodegenData.put {intros = intros, graph = graph,
@@ -188,7 +188,7 @@
         end)) (AList.lookup op = modes name)
 
   in (case strip_comb t of
-      (Const (@{const_name "op ="}, Type (_, [T, _])), _) =>
+      (Const (@{const_name HOL.eq}, Type (_, [T, _])), _) =>
         [Mode ((([], [1]), false), [1], []), Mode ((([], [2]), false), [2], [])] @
         (if is_eqT T then [Mode ((([], [1, 2]), false), [1, 2], [])] else [])
     | (Const (name, _), args) => the_default default (mk_modes name args)
@@ -344,7 +344,7 @@
   end;
 
 fun modename module s (iss, is) gr =
-  let val (id, gr') = if s = @{const_name "op ="} then (("", "equal"), gr)
+  let val (id, gr') = if s = @{const_name HOL.eq} then (("", "equal"), gr)
     else mk_const_id module s gr
   in (space_implode "__"
     (mk_qual_id module id ::
@@ -363,7 +363,7 @@
   | compile_expr thy defs dep module brack modes (SOME (Mode ((mode, _), _, ms)), t) gr =
       (case strip_comb t of
          (Const (name, _), args) =>
-           if name = @{const_name "op ="} orelse AList.defined op = modes name then
+           if name = @{const_name HOL.eq} orelse AList.defined op = modes name then
              let
                val (args1, args2) = chop (length ms) args;
                val ((ps, mode_id), gr') = gr |> fold_map
@@ -581,7 +581,7 @@
 
       fun dest_prem (_ $ (Const (@{const_name Set.member}, _) $ t $ u)) =
             Prem ([t], Envir.beta_eta_contract u, true)
-        | dest_prem (_ $ ((eq as Const (@{const_name "op ="}, _)) $ t $ u)) =
+        | dest_prem (_ $ ((eq as Const (@{const_name HOL.eq}, _)) $ t $ u)) =
             Prem ([t, u], eq, false)
         | dest_prem (_ $ t) =
             (case strip_comb t of
--- a/src/HOL/Tools/inductive_set.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/inductive_set.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -74,9 +74,9 @@
         in Drule.instantiate' [] (rev (map (SOME o cterm_of thy o Var) vs))
           (p (fold (Logic.all o Var) vs t) f)
         end;
-      fun mkop @{const_name "op &"} T x =
+      fun mkop @{const_name HOL.conj} T x =
             SOME (Const (@{const_name Lattices.inf}, T --> T --> T), x)
-        | mkop @{const_name "op |"} T x =
+        | mkop @{const_name HOL.disj} T x =
             SOME (Const (@{const_name Lattices.sup}, T --> T --> T), x)
         | mkop _ _ _ = NONE;
       fun mk_collect p T t =
@@ -265,7 +265,7 @@
 
 fun add ctxt thm (tab as {to_set_simps, to_pred_simps, set_arities, pred_arities}) =
   case prop_of thm of
-    Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, Type (_, [T, _])) $ lhs $ rhs) =>
+    Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, Type (_, [T, _])) $ lhs $ rhs) =>
       (case body_type T of
          @{typ bool} =>
            let
--- a/src/HOL/Tools/int_arith.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/int_arith.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -51,7 +51,7 @@
 
 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 ="},
+  | check (Const (s, _)) = member (op =) [@{const_name HOL.eq},
       @{const_name Groups.times}, @{const_name Groups.uminus},
       @{const_name Groups.minus}, @{const_name Groups.plus},
       @{const_name Groups.zero},
@@ -91,7 +91,7 @@
 fun number_of thy T n =
   if not (Sign.of_sort thy (T, @{sort number}))
   then raise CTERM ("number_of", [])
-  else Numeral.mk_cnumber (Thm.ctyp_of thy T) n
+  else Numeral.mk_cnumber (Thm.ctyp_of thy T) n;
 
 val setup =
   Lin_Arith.add_inj_thms [@{thm zle_int} RS iffD2, @{thm int_int_eq} RS iffD2]
--- a/src/HOL/Tools/lin_arith.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/lin_arith.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -16,8 +16,7 @@
   val add_simprocs: simproc list -> Context.generic -> Context.generic
   val add_inj_const: string * typ -> Context.generic -> Context.generic
   val add_discrete_type: string -> Context.generic -> Context.generic
-  val set_number_of: (theory -> typ -> int -> cterm) -> Context.generic ->
-    Context.generic
+  val set_number_of: (theory -> typ -> int -> cterm) -> Context.generic -> Context.generic
   val setup: Context.generic -> Context.generic
   val global_setup: theory -> theory
   val split_limit: int Config.T
@@ -46,7 +45,7 @@
 val mk_Trueprop = HOLogic.mk_Trueprop;
 
 fun atomize thm = case Thm.prop_of thm of
-    Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op &"}, _) $ _ $ _) =>
+    Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.conj}, _) $ _ $ _) =>
     atomize (thm RS conjunct1) @ atomize (thm RS conjunct2)
   | _ => [thm];
 
@@ -230,7 +229,7 @@
   case rel of
     @{const_name Orderings.less}    => SOME (p, i, "<", q, j)
   | @{const_name Orderings.less_eq} => SOME (p, i, "<=", q, j)
-  | "op ="              => SOME (p, i, "=", q, j)
+  | @{const_name HOL.eq}            => SOME (p, i, "=", q, j)
   | _                   => NONE
 end handle Rat.DIVZERO => NONE;
 
@@ -428,7 +427,7 @@
         val t2'             = incr_boundvars 1 t2
         val t1_lt_t2        = Const (@{const_name Orderings.less},
                                 split_type --> split_type --> HOLogic.boolT) $ t1 $ t2
-        val t1_eq_t2_plus_d = Const (@{const_name "op ="}, split_type --> split_type --> HOLogic.boolT) $ t1' $
+        val t1_eq_t2_plus_d = Const (@{const_name HOL.eq}, split_type --> split_type --> HOLogic.boolT) $ t1' $
                                 (Const (@{const_name Groups.plus},
                                   split_type --> split_type --> split_type) $ t2' $ d)
         val not_false       = HOLogic.mk_Trueprop (HOLogic.Not $ HOLogic.false_const)
@@ -448,7 +447,7 @@
                             (map (incr_boundvars 1) rev_terms)
         val terms2      = map (subst_term [(split_term, zero_nat)]) rev_terms
         val t1'         = incr_boundvars 1 t1
-        val t1_eq_nat_n = Const (@{const_name "op ="}, HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ t1' $
+        val t1_eq_nat_n = Const (@{const_name HOL.eq}, HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ t1' $
                             (Const (@{const_name of_nat}, HOLogic.natT --> HOLogic.intT) $ n)
         val t1_lt_zero  = Const (@{const_name Orderings.less},
                             HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ t1 $ zero_int
@@ -472,13 +471,13 @@
                                         (map (incr_boundvars 2) rev_terms)
         val t1'                     = incr_boundvars 2 t1
         val t2'                     = incr_boundvars 2 t2
-        val t2_eq_zero              = Const (@{const_name "op ="},
+        val t2_eq_zero              = Const (@{const_name HOL.eq},
                                         split_type --> split_type --> HOLogic.boolT) $ t2 $ zero
-        val t2_neq_zero             = HOLogic.mk_not (Const (@{const_name "op ="},
+        val t2_neq_zero             = HOLogic.mk_not (Const (@{const_name HOL.eq},
                                         split_type --> split_type --> HOLogic.boolT) $ t2' $ zero)
         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 (@{const_name "op ="}, split_type --> split_type --> HOLogic.boolT) $ t1' $
+        val t1_eq_t2_times_i_plus_j = Const (@{const_name HOL.eq}, split_type --> split_type --> HOLogic.boolT) $ t1' $
                                        (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)
@@ -504,13 +503,13 @@
                                         (map (incr_boundvars 2) rev_terms)
         val t1'                     = incr_boundvars 2 t1
         val t2'                     = incr_boundvars 2 t2
-        val t2_eq_zero              = Const (@{const_name "op ="},
+        val t2_eq_zero              = Const (@{const_name HOL.eq},
                                         split_type --> split_type --> HOLogic.boolT) $ t2 $ zero
-        val t2_neq_zero             = HOLogic.mk_not (Const (@{const_name "op ="},
+        val t2_neq_zero             = HOLogic.mk_not (Const (@{const_name HOL.eq},
                                         split_type --> split_type --> HOLogic.boolT) $ t2' $ zero)
         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 (@{const_name "op ="}, split_type --> split_type --> HOLogic.boolT) $ t1' $
+        val t1_eq_t2_times_i_plus_j = Const (@{const_name HOL.eq}, split_type --> split_type --> HOLogic.boolT) $ t1' $
                                        (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)
@@ -542,7 +541,7 @@
                                         (map (incr_boundvars 2) rev_terms)
         val t1'                     = incr_boundvars 2 t1
         val t2'                     = incr_boundvars 2 t2
-        val t2_eq_zero              = Const (@{const_name "op ="},
+        val t2_eq_zero              = Const (@{const_name HOL.eq},
                                         split_type --> split_type --> HOLogic.boolT) $ t2 $ zero
         val zero_lt_t2              = Const (@{const_name Orderings.less},
                                         split_type --> split_type --> HOLogic.boolT) $ zero $ t2'
@@ -556,7 +555,7 @@
                                         split_type --> split_type--> HOLogic.boolT) $ j $ t2'
         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 (@{const_name "op ="}, split_type --> split_type --> HOLogic.boolT) $ t1' $
+        val t1_eq_t2_times_i_plus_j = Const (@{const_name HOL.eq}, split_type --> split_type --> HOLogic.boolT) $ t1' $
                                        (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)
@@ -596,7 +595,7 @@
                                         (map (incr_boundvars 2) rev_terms)
         val t1'                     = incr_boundvars 2 t1
         val t2'                     = incr_boundvars 2 t2
-        val t2_eq_zero              = Const (@{const_name "op ="},
+        val t2_eq_zero              = Const (@{const_name HOL.eq},
                                         split_type --> split_type --> HOLogic.boolT) $ t2 $ zero
         val zero_lt_t2              = Const (@{const_name Orderings.less},
                                         split_type --> split_type --> HOLogic.boolT) $ zero $ t2'
@@ -610,7 +609,7 @@
                                         split_type --> split_type--> HOLogic.boolT) $ j $ t2'
         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 (@{const_name "op ="}, split_type --> split_type --> HOLogic.boolT) $ t1' $
+        val t1_eq_t2_times_i_plus_j = Const (@{const_name HOL.eq}, split_type --> split_type --> HOLogic.boolT) $ t1' $
                                        (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)
@@ -769,29 +768,11 @@
 
 structure Fast_Arith = Fast_Lin_Arith(structure LA_Logic = LA_Logic and LA_Data = LA_Data);
 
-fun map_inj_thms f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
-  {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = f inj_thms,
-    lessD = lessD, neqE = neqE, simpset = simpset, number_of = number_of};
-
-fun map_lessD f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
-  {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms,
-    lessD = f lessD, neqE = neqE, simpset = simpset, number_of = number_of};
-
-fun map_simpset f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
-  {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms,
-    lessD = lessD, neqE = neqE, simpset = f simpset, number_of = number_of};
-
-fun map_number_of f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
-  {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms,
-    lessD = lessD, neqE = neqE, simpset = simpset, number_of = f number_of};
-
-fun add_inj_thms thms = Fast_Arith.map_data (map_inj_thms (append thms));
-fun add_lessD thm = Fast_Arith.map_data (map_lessD (fn thms => thms @ [thm]));
-fun add_simps thms = Fast_Arith.map_data (map_simpset (fn simpset => simpset addsimps thms));
-fun add_simprocs procs = Fast_Arith.map_data (map_simpset (fn simpset => simpset addsimprocs procs));
-
-fun set_number_of f = Fast_Arith.map_data (map_number_of (K (serial (), f)))
-
+val add_inj_thms = Fast_Arith.add_inj_thms;
+val add_lessD = Fast_Arith.add_lessD;
+val add_simps = Fast_Arith.add_simps;
+val add_simprocs = Fast_Arith.add_simprocs;
+val set_number_of = Fast_Arith.set_number_of;
 
 fun simple_tac ctxt = Fast_Arith.lin_arith_tac ctxt false;
 val lin_arith_tac = Fast_Arith.lin_arith_tac;
--- a/src/HOL/Tools/list_code.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/list_code.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -46,7 +46,7 @@
             Code_Printer.literal_list literals (map (pr vars Code_Printer.NOBR) ts)
         | NONE =>
             default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;
-  in Code_Target.add_syntax_const target @{const_name Cons}
+  in Code_Target.add_const_syntax target @{const_name Cons}
     (SOME (Code_Printer.complex_const_syntax (2, ([@{const_name Nil}, @{const_name Cons}], pretty))))
   end
 
--- a/src/HOL/Tools/meson.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/meson.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -48,7 +48,10 @@
 fun trace_msg msg = if ! trace then tracing (msg ()) else ();
 
 val max_clauses_default = 60;
-val (max_clauses, setup) = Attrib.config_int "max_clauses" (K max_clauses_default);
+val (max_clauses, setup) = Attrib.config_int "meson_max_clauses" (K max_clauses_default);
+
+(*No known example (on 1-5-2007) needs even thirty*)
+val iter_deepen_limit = 50;
 
 val disj_forward = @{thm disj_forward};
 val disj_forward2 = @{thm disj_forward2};
@@ -151,8 +154,8 @@
   let fun has (Const(a,_)) = false
         | has (Const(@{const_name Trueprop},_) $ p) = has p
         | has (Const(@{const_name Not},_) $ p) = has p
-        | has (Const(@{const_name "op |"},_) $ p $ q) = member (op =) bs @{const_name "op |"} orelse has p orelse has q
-        | has (Const(@{const_name "op &"},_) $ p $ q) = member (op =) bs @{const_name "op &"} orelse has p orelse has q
+        | has (Const(@{const_name HOL.disj},_) $ p $ q) = member (op =) bs @{const_name HOL.disj} orelse has p orelse has q
+        | has (Const(@{const_name HOL.conj},_) $ p $ q) = member (op =) bs @{const_name HOL.conj} orelse has p orelse has q
         | has (Const(@{const_name All},_) $ Abs(_,_,p)) = member (op =) bs @{const_name All} orelse has p
         | has (Const(@{const_name Ex},_) $ Abs(_,_,p)) = member (op =) bs @{const_name Ex} orelse has p
         | has _ = false
@@ -162,7 +165,7 @@
 (**** Clause handling ****)
 
 fun literals (Const(@{const_name Trueprop},_) $ P) = literals P
-  | literals (Const(@{const_name "op |"},_) $ P $ Q) = literals P @ literals Q
+  | literals (Const(@{const_name HOL.disj},_) $ P $ Q) = literals P @ literals Q
   | literals (Const(@{const_name Not},_) $ P) = [(false,P)]
   | literals P = [(true,P)];
 
@@ -172,7 +175,7 @@
 
 (*** Tautology Checking ***)
 
-fun signed_lits_aux (Const (@{const_name "op |"}, _) $ P $ Q) (poslits, neglits) =
+fun signed_lits_aux (Const (@{const_name HOL.disj}, _) $ P $ Q) (poslits, neglits) =
       signed_lits_aux Q (signed_lits_aux P (poslits, neglits))
   | signed_lits_aux (Const(@{const_name Not},_) $ P) (poslits, neglits) = (poslits, P::neglits)
   | signed_lits_aux P (poslits, neglits) = (P::poslits, neglits);
@@ -180,7 +183,7 @@
 fun signed_lits th = signed_lits_aux (HOLogic.dest_Trueprop (concl_of th)) ([],[]);
 
 (*Literals like X=X are tautologous*)
-fun taut_poslit (Const(@{const_name "op ="},_) $ t $ u) = t aconv u
+fun taut_poslit (Const(@{const_name HOL.eq},_) $ t $ u) = t aconv u
   | taut_poslit (Const(@{const_name True},_)) = true
   | taut_poslit _ = false;
 
@@ -208,18 +211,18 @@
 fun refl_clause_aux 0 th = th
   | refl_clause_aux n th =
        case HOLogic.dest_Trueprop (concl_of th) of
-          (Const (@{const_name "op |"}, _) $ (Const (@{const_name "op |"}, _) $ _ $ _) $ _) =>
+          (Const (@{const_name HOL.disj}, _) $ (Const (@{const_name HOL.disj}, _) $ _ $ _) $ _) =>
             refl_clause_aux n (th RS disj_assoc)    (*isolate an atom as first disjunct*)
-        | (Const (@{const_name "op |"}, _) $ (Const(@{const_name Not},_) $ (Const(@{const_name "op ="},_) $ t $ u)) $ _) =>
+        | (Const (@{const_name HOL.disj}, _) $ (Const(@{const_name Not},_) $ (Const(@{const_name HOL.eq},_) $ t $ u)) $ _) =>
             if eliminable(t,u)
             then refl_clause_aux (n-1) (th RS not_refl_disj_D)  (*Var inequation: delete*)
             else refl_clause_aux (n-1) (th RS disj_comm)  (*not between Vars: ignore*)
-        | (Const (@{const_name "op |"}, _) $ _ $ _) => refl_clause_aux n (th RS disj_comm)
+        | (Const (@{const_name HOL.disj}, _) $ _ $ _) => refl_clause_aux n (th RS disj_comm)
         | _ => (*not a disjunction*) th;
 
-fun notequal_lits_count (Const (@{const_name "op |"}, _) $ P $ Q) =
+fun notequal_lits_count (Const (@{const_name HOL.disj}, _) $ P $ Q) =
       notequal_lits_count P + notequal_lits_count Q
-  | notequal_lits_count (Const(@{const_name Not},_) $ (Const(@{const_name "op ="},_) $ _ $ _)) = 1
+  | notequal_lits_count (Const(@{const_name Not},_) $ (Const(@{const_name HOL.eq},_) $ _ $ _)) = 1
   | notequal_lits_count _ = 0;
 
 (*Simplify a clause by applying reflexivity to its negated equality literals*)
@@ -268,16 +271,16 @@
   (*Estimate the number of clauses in order to detect infeasible theorems*)
   fun signed_nclauses b (Const(@{const_name Trueprop},_) $ t) = signed_nclauses b t
     | signed_nclauses b (Const(@{const_name Not},_) $ t) = signed_nclauses (not b) t
-    | signed_nclauses b (Const(@{const_name "op &"},_) $ t $ u) =
+    | signed_nclauses b (Const(@{const_name HOL.conj},_) $ t $ u) =
         if b then sum (signed_nclauses b t) (signed_nclauses b u)
              else prod (signed_nclauses b t) (signed_nclauses b u)
-    | signed_nclauses b (Const(@{const_name "op |"},_) $ t $ u) =
+    | signed_nclauses b (Const(@{const_name HOL.disj},_) $ t $ u) =
         if b then prod (signed_nclauses b t) (signed_nclauses b u)
              else sum (signed_nclauses b t) (signed_nclauses b u)
-    | signed_nclauses b (Const(@{const_name "op -->"},_) $ t $ u) =
+    | signed_nclauses b (Const(@{const_name HOL.implies},_) $ t $ u) =
         if b then prod (signed_nclauses (not b) t) (signed_nclauses b u)
              else sum (signed_nclauses (not b) t) (signed_nclauses b u)
-    | signed_nclauses b (Const(@{const_name "op ="}, Type ("fun", [T, _])) $ t $ u) =
+    | signed_nclauses b (Const(@{const_name HOL.eq}, Type ("fun", [T, _])) $ t $ u) =
         if T = HOLogic.boolT then (*Boolean equality is if-and-only-if*)
             if b then sum (prod (signed_nclauses (not b) t) (signed_nclauses b u))
                           (prod (signed_nclauses (not b) u) (signed_nclauses b t))
@@ -330,10 +333,10 @@
   let val ctxtr = Unsynchronized.ref ctxt   (* FIXME ??? *)
       fun cnf_aux (th,ths) =
         if not (can HOLogic.dest_Trueprop (prop_of th)) then ths (*meta-level: ignore*)
-        else if not (has_conns [@{const_name All}, @{const_name Ex}, @{const_name "op &"}] (prop_of th))
+        else if not (has_conns [@{const_name All}, @{const_name Ex}, @{const_name HOL.conj}] (prop_of th))
         then nodups ctxt th :: ths (*no work to do, terminate*)
         else case head_of (HOLogic.dest_Trueprop (concl_of th)) of
-            Const (@{const_name "op &"}, _) => (*conjunction*)
+            Const (@{const_name HOL.conj}, _) => (*conjunction*)
                 cnf_aux (th RS conjunct1, cnf_aux (th RS conjunct2, ths))
           | Const (@{const_name All}, _) => (*universal quantifier*)
                 let val (th',ctxt') = freeze_spec th (!ctxtr)
@@ -341,7 +344,7 @@
           | Const (@{const_name Ex}, _) =>
               (*existential quantifier: Insert Skolem functions*)
               cnf_aux (apply_skolem_theorem (th, skolem_ths), ths)
-          | Const (@{const_name "op |"}, _) =>
+          | Const (@{const_name HOL.disj}, _) =>
               (*Disjunction of P, Q: Create new goal of proving ?P | ?Q and solve it using
                 all combinations of converting P, Q to CNF.*)
               let val tac =
@@ -365,7 +368,7 @@
 (**** Generation of contrapositives ****)
 
 fun is_left (Const (@{const_name Trueprop}, _) $
-               (Const (@{const_name "op |"}, _) $ (Const (@{const_name "op |"}, _) $ _ $ _) $ _)) = true
+               (Const (@{const_name HOL.disj}, _) $ (Const (@{const_name HOL.disj}, _) $ _ $ _) $ _)) = true
   | is_left _ = false;
 
 (*Associate disjuctions to right -- make leftmost disjunct a LITERAL*)
@@ -400,8 +403,8 @@
 (*Is the string the name of a connective? Really only | and Not can remain,
   since this code expects to be called on a clause form.*)
 val is_conn = member (op =)
-    [@{const_name Trueprop}, @{const_name "op &"}, @{const_name "op |"},
-     @{const_name "op -->"}, @{const_name Not},
+    [@{const_name Trueprop}, @{const_name HOL.conj}, @{const_name HOL.disj},
+     @{const_name HOL.implies}, @{const_name Not},
      @{const_name All}, @{const_name Ex}, @{const_name Ball}, @{const_name Bex}];
 
 (*True if the term contains a function--not a logical connective--where the type
@@ -429,7 +432,7 @@
 
 fun rigid t = not (is_Var (head_of t));
 
-fun ok4horn (Const (@{const_name Trueprop},_) $ (Const (@{const_name "op |"}, _) $ t $ _)) = rigid t
+fun ok4horn (Const (@{const_name Trueprop},_) $ (Const (@{const_name HOL.disj}, _) $ t $ _)) = rigid t
   | ok4horn (Const (@{const_name Trueprop},_) $ t) = rigid t
   | ok4horn _ = false;
 
@@ -640,7 +643,7 @@
     end;
 
 fun iter_deepen_prolog_tac horns =
-    ITER_DEEPEN (has_fewer_prems 1) (prolog_step_tac' horns);
+    ITER_DEEPEN iter_deepen_limit (has_fewer_prems 1) (prolog_step_tac' horns);
 
 fun iter_deepen_meson_tac ctxt ths = ctxt |> MESON make_clauses
   (fn cls =>
@@ -653,7 +656,10 @@
             cat_lines ("meson method called:" ::
               map (Display.string_of_thm ctxt) (cls @ ths) @
               ["clauses:"] @ map (Display.string_of_thm ctxt) horns))
-        in THEN_ITER_DEEPEN (resolve_tac goes 1) (has_fewer_prems 1) (prolog_step_tac' horns) end));
+        in
+          THEN_ITER_DEEPEN iter_deepen_limit
+            (resolve_tac goes 1) (has_fewer_prems 1) (prolog_step_tac' horns)
+        end));
 
 fun meson_tac ctxt ths =
   SELECT_GOAL (TRY (safe_tac (claset_of ctxt)) THEN TRYALL (iter_deepen_meson_tac ctxt ths));
--- a/src/HOL/Tools/nat_arith.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/nat_arith.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -62,7 +62,7 @@
 (struct
   open CommonCancelSums;
   val mk_bal = HOLogic.mk_eq;
-  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT;
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} HOLogic.natT;
   val uncancel_tac = gen_uncancel_tac @{thm "nat_add_left_cancel"};
 end);
 
--- a/src/HOL/Tools/nat_numeral_simprocs.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/nat_numeral_simprocs.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -168,7 +168,7 @@
  (open CancelNumeralsCommon
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} HOLogic.natT
   val bal_add1 = @{thm nat_eq_add_iff1} RS trans
   val bal_add2 = @{thm nat_eq_add_iff2} RS trans
 );
@@ -300,7 +300,7 @@
  (open CancelNumeralFactorCommon
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} HOLogic.natT
   val cancel = @{thm nat_mult_eq_cancel1} RS trans
   val neg_exchanges = false
 )
@@ -380,7 +380,7 @@
  (open CancelFactorCommon
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} HOLogic.natT
   fun simp_conv _ _ = SOME @{thm nat_mult_eq_cancel_disj}
 );
 
--- a/src/HOL/Tools/numeral.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/numeral.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -76,7 +76,7 @@
     fun pretty literals [pls', min', bit0', bit1'] _ thm _ _ [(t, _)] =
       (Code_Printer.str o print literals o the_default 0 o dest_numeral pls' min' bit0' bit1' thm) t;
   in
-    thy |> Code_Target.add_syntax_const target number_of
+    thy |> Code_Target.add_const_syntax target number_of
       (SOME (Code_Printer.complex_const_syntax (1, ([@{const_name Int.Pls}, @{const_name Int.Min},
         @{const_name Int.Bit0}, @{const_name Int.Bit1}], pretty))))
   end;
--- a/src/HOL/Tools/numeral_simprocs.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/numeral_simprocs.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -222,7 +222,7 @@
  (open CancelNumeralsCommon
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} Term.dummyT
   val bal_add1 = @{thm eq_add_iff1} RS trans
   val bal_add2 = @{thm eq_add_iff2} RS trans
 );
@@ -401,7 +401,7 @@
  (open CancelNumeralFactorCommon
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} Term.dummyT
   val cancel = @{thm mult_cancel_left} RS trans
   val neg_exchanges = false
 )
@@ -516,7 +516,7 @@
  (open CancelFactorCommon
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} Term.dummyT
   fun simp_conv _ _ = SOME @{thm mult_cancel_left}
 );
 
@@ -606,7 +606,7 @@
 local
  val zr = @{cpat "0"}
  val zT = ctyp_of_term zr
- val geq = @{cpat "op ="}
+ val geq = @{cpat HOL.eq}
  val eqT = Thm.dest_ctyp (ctyp_of_term geq) |> hd
  val add_frac_eq = mk_meta_eq @{thm "add_frac_eq"}
  val add_frac_num = mk_meta_eq @{thm "add_frac_num"}
@@ -676,7 +676,7 @@
         val T = ctyp_of_term c
         val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_le_eq"}
       in SOME (mk_meta_eq th) end
-  | Const(@{const_name "op ="},_)$(Const(@{const_name Rings.divide},_)$_$_)$_ =>
+  | Const(@{const_name HOL.eq},_)$(Const(@{const_name Rings.divide},_)$_$_)$_ =>
       let
         val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop
         val _ = map is_number [a,b,c]
@@ -697,7 +697,7 @@
         val T = ctyp_of_term c
         val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "le_divide_eq"}
       in SOME (mk_meta_eq th) end
-  | Const(@{const_name "op ="},_)$_$(Const(@{const_name Rings.divide},_)$_$_) =>
+  | Const(@{const_name HOL.eq},_)$_$(Const(@{const_name Rings.divide},_)$_$_) =>
     let
       val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop
         val _ = map is_number [a,b,c]
--- a/src/HOL/Tools/prop_logic.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/prop_logic.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -397,14 +397,14 @@
 			(False, table)
 		  | aux (Const (@{const_name Not}, _) $ x)      table =
 			apfst Not (aux x table)
-		  | aux (Const (@{const_name "op |"}, _) $ x $ y) table =
+		  | aux (Const (@{const_name HOL.disj}, _) $ x $ y) table =
 			let
 				val (fm1, table1) = aux x table
 				val (fm2, table2) = aux y table1
 			in
 				(Or (fm1, fm2), table2)
 			end
-		  | aux (Const (@{const_name "op &"}, _) $ x $ y) table =
+		  | aux (Const (@{const_name HOL.conj}, _) $ x $ y) table =
 			let
 				val (fm1, table1) = aux x table
 				val (fm2, table2) = aux y table1
--- a/src/HOL/Tools/quickcheck_generators.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/quickcheck_generators.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -342,7 +342,7 @@
     val bound_max = length Ts - 1;
     val bounds = map_index (fn (i, ty) =>
       (2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) Ts;
-    fun strip_imp (Const(@{const_name "op -->"},_) $ A $ B) = apfst (cons A) (strip_imp B)
+    fun strip_imp (Const(@{const_name HOL.implies},_) $ A $ B) = apfst (cons A) (strip_imp B)
       | strip_imp A = ([], A)
     val prop' = betapplys (prop, map (fn (i, _, _, _) => Bound i) bounds);
     val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) bounds)
--- a/src/HOL/Tools/recfun_codegen.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/recfun_codegen.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -40,7 +40,7 @@
       end
   | avoid_value thy thms = thms;
 
-fun get_equations thy defs (raw_c, T) = if raw_c = @{const_name "op ="} then ([], "") else
+fun get_equations thy defs (raw_c, T) = if raw_c = @{const_name HOL.eq} then ([], "") else
   let
     val c = AxClass.unoverload_const thy (raw_c, T);
     val raw_thms = Code.get_cert thy (Code_Preproc.preprocess_functrans thy) c
--- a/src/HOL/Tools/record.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/record.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -420,7 +420,7 @@
   equalities = equalities, extinjects=extinjects, extsplit = extsplit, splits = splits,
   extfields = extfields, fieldext = fieldext }: data;
 
-structure Records_Data = Theory_Data
+structure Data = Theory_Data
 (
   type T = data;
   val empty =
@@ -474,25 +474,22 @@
 
 (* access 'records' *)
 
-val get_info = Symtab.lookup o #records o Records_Data.get;
+val get_info = Symtab.lookup o #records o Data.get;
 
 fun the_info thy name =
   (case get_info thy name of
     SOME info => info
   | NONE => error ("Unknown record type " ^ quote name));
 
-fun put_record name info thy =
-  let
-    val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
-      Records_Data.get thy;
-    val data = make_data (Symtab.update (name, info) records)
-      sel_upd equalities extinjects extsplit splits extfields fieldext;
-  in Records_Data.put data thy end;
+fun put_record name info =
+  Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+    make_data (Symtab.update (name, info) records)
+      sel_upd equalities extinjects extsplit splits extfields fieldext);
 
 
 (* access 'sel_upd' *)
 
-val get_sel_upd = #sel_upd o Records_Data.get;
+val get_sel_upd = #sel_upd o Data.get;
 
 val is_selector = Symtab.defined o #selectors o get_sel_upd;
 val get_updates = Symtab.lookup o #updates o get_sel_upd;
@@ -517,7 +514,7 @@
     val upds = map (suffix updateN) all ~~ all;
 
     val {records, sel_upd = {selectors, updates, simpset, defset, foldcong, unfoldcong},
-      equalities, extinjects, extsplit, splits, extfields, fieldext} = Records_Data.get thy;
+      equalities, extinjects, extsplit, splits, extfields, fieldext} = Data.get thy;
     val data = make_data records
       {selectors = fold Symtab.update_new sels selectors,
         updates = fold Symtab.update_new upds updates,
@@ -526,80 +523,61 @@
         foldcong = foldcong addcongs folds,
         unfoldcong = unfoldcong addcongs unfolds}
        equalities extinjects extsplit splits extfields fieldext;
-  in Records_Data.put data thy end;
+  in Data.put data thy end;
 
 
 (* access 'equalities' *)
 
-fun add_equalities name thm thy =
-  let
-    val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
-      Records_Data.get thy;
-    val data = make_data records sel_upd
-      (Symtab.update_new (name, thm) equalities) extinjects extsplit splits extfields fieldext;
-  in Records_Data.put data thy end;
-
-val get_equalities = Symtab.lookup o #equalities o Records_Data.get;
+fun add_equalities name thm =
+  Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+    make_data records sel_upd
+      (Symtab.update_new (name, thm) equalities) extinjects extsplit splits extfields fieldext);
+
+val get_equalities = Symtab.lookup o #equalities o Data.get;
 
 
 (* access 'extinjects' *)
 
-fun add_extinjects thm thy =
-  let
-    val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
-      Records_Data.get thy;
-    val data =
-      make_data records sel_upd equalities (insert Thm.eq_thm_prop thm extinjects)
-        extsplit splits extfields fieldext;
-  in Records_Data.put data thy end;
-
-val get_extinjects = rev o #extinjects o Records_Data.get;
+fun add_extinjects thm =
+  Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+    make_data records sel_upd equalities (insert Thm.eq_thm_prop thm extinjects)
+      extsplit splits extfields fieldext);
+
+val get_extinjects = rev o #extinjects o Data.get;
 
 
 (* access 'extsplit' *)
 
-fun add_extsplit name thm thy =
-  let
-    val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
-      Records_Data.get thy;
-    val data =
-      make_data records sel_upd equalities extinjects
-        (Symtab.update_new (name, thm) extsplit) splits extfields fieldext;
-  in Records_Data.put data thy end;
+fun add_extsplit name thm =
+  Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+    make_data records sel_upd equalities extinjects
+      (Symtab.update_new (name, thm) extsplit) splits extfields fieldext);
 
 
 (* access 'splits' *)
 
-fun add_splits name thmP thy =
-  let
-    val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
-      Records_Data.get thy;
-    val data =
-      make_data records sel_upd equalities extinjects extsplit
-        (Symtab.update_new (name, thmP) splits) extfields fieldext;
-  in Records_Data.put data thy end;
-
-val get_splits = Symtab.lookup o #splits o Records_Data.get;
+fun add_splits name thmP =
+  Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+    make_data records sel_upd equalities extinjects extsplit
+      (Symtab.update_new (name, thmP) splits) extfields fieldext);
+
+val get_splits = Symtab.lookup o #splits o Data.get;
 
 
 (* parent/extension of named record *)
 
-val get_parent = (Option.join o Option.map #parent) oo (Symtab.lookup o #records o Records_Data.get);
-val get_extension = Option.map #extension oo (Symtab.lookup o #records o Records_Data.get);
+val get_parent = (Option.join o Option.map #parent) oo (Symtab.lookup o #records o Data.get);
+val get_extension = Option.map #extension oo (Symtab.lookup o #records o Data.get);
 
 
 (* access 'extfields' *)
 
-fun add_extfields name fields thy =
-  let
-    val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
-      Records_Data.get thy;
-    val data =
-      make_data records sel_upd equalities extinjects extsplit splits
-        (Symtab.update_new (name, fields) extfields) fieldext;
-  in Records_Data.put data thy end;
-
-val get_extfields = Symtab.lookup o #extfields o Records_Data.get;
+fun add_extfields name fields =
+  Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+    make_data records sel_upd equalities extinjects extsplit splits
+      (Symtab.update_new (name, fields) extfields) fieldext);
+
+val get_extfields = Symtab.lookup o #extfields o Data.get;
 
 fun get_extT_fields thy T =
   let
@@ -609,7 +587,7 @@
       in Long_Name.implode (rev (nm :: rst)) end;
     val midx = maxidx_of_typs (moreT :: Ts);
     val varifyT = varifyT midx;
-    val {records, extfields, ...} = Records_Data.get thy;
+    val {records, extfields, ...} = Data.get thy;
     val (fields, (more, _)) = split_last (Symtab.lookup_list extfields name);
     val args = map varifyT (snd (#extension (the (Symtab.lookup records recname))));
 
@@ -628,18 +606,14 @@
 
 (* access 'fieldext' *)
 
-fun add_fieldext extname_types fields thy =
-  let
-    val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
-      Records_Data.get thy;
-    val fieldext' =
-      fold (fn field => Symtab.update_new (field, extname_types)) fields fieldext;
-    val data =
-      make_data records sel_upd equalities extinjects
-        extsplit splits extfields fieldext';
-  in Records_Data.put data thy end;
-
-val get_fieldext = Symtab.lookup o #fieldext o Records_Data.get;
+fun add_fieldext extname_types fields =
+  Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+    let
+      val fieldext' =
+        fold (fn field => Symtab.update_new (field, extname_types)) fields fieldext;
+    in make_data records sel_upd equalities extinjects extsplit splits extfields fieldext' end);
+
+val get_fieldext = Symtab.lookup o #fieldext o Data.get;
 
 
 (* parent records *)
@@ -1179,7 +1153,7 @@
             ((upd as Const (u, Type (_, [_, Type (_, [rT, _])]))) $ k $ r) =>
           if is_selector thy s andalso is_some (get_updates thy u) then
             let
-              val {sel_upd = {updates, ...}, extfields, ...} = Records_Data.get thy;
+              val {sel_upd = {updates, ...}, extfields, ...} = Data.get thy;
 
               fun mk_eq_terms ((upd as Const (u, Type(_, [kT, _]))) $ k $ r) =
                     (case Symtab.lookup updates u of
@@ -1368,7 +1342,7 @@
 val eq_simproc =
   Simplifier.simproc_global @{theory HOL} "record_eq_simp" ["r = s"]
     (fn thy => fn _ => fn t =>
-      (case t of Const (@{const_name "op ="}, Type (_, [T, _])) $ _ $ _ =>
+      (case t of Const (@{const_name HOL.eq}, Type (_, [T, _])) $ _ $ _ =>
         (case rec_id ~1 T of
           "" => NONE
         | name =>
@@ -1431,12 +1405,12 @@
                 else raise TERM ("", [x]);
               val sel' = Const (sel, Tsel) $ Bound 0;
               val (l, r) = if lr then (sel', x') else (x', sel');
-            in Const (@{const_name "op ="}, Teq) $ l $ r end
+            in Const (@{const_name HOL.eq}, Teq) $ l $ r end
           else raise TERM ("", [Const (sel, Tsel)]);
 
-        fun dest_sel_eq (Const (@{const_name "op ="}, Teq) $ (Const (sel, Tsel) $ Bound 0) $ X) =
+        fun dest_sel_eq (Const (@{const_name HOL.eq}, Teq) $ (Const (sel, Tsel) $ Bound 0) $ X) =
               (true, Teq, (sel, Tsel), X)
-          | dest_sel_eq (Const (@{const_name "op ="}, Teq) $ X $ (Const (sel, Tsel) $ Bound 0)) =
+          | dest_sel_eq (Const (@{const_name HOL.eq}, Teq) $ X $ (Const (sel, Tsel) $ Bound 0)) =
               (false, Teq, (sel, Tsel), X)
           | dest_sel_eq _ = raise TERM ("", []);
       in
@@ -1598,15 +1572,17 @@
         [] => (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop
           (hd (rev (Logic.strip_assums_hyp (hd (prems_of rule')))))))), true)
       | [x] => (head_of x, false));
-    val rule'' = cterm_instantiate (map (pairself cert)
-      (case rev params of
-        [] =>
-          (case AList.lookup (op =) (Term.add_frees g []) s of
-            NONE => sys_error "try_param_tac: no such variable"
-          | SOME T => [(P, if ca then concl else lambda (Free (s, T)) concl), (x, Free (s, T))])
-      | (_, T) :: _ =>
-          [(P, list_abs (params, if ca then concl else incr_boundvars 1 (Abs (s, T, concl)))),
-            (x, list_abs (params, Bound 0))])) rule';
+    val rule'' =
+      cterm_instantiate
+        (map (pairself cert)
+          (case rev params of
+            [] =>
+              (case AList.lookup (op =) (Term.add_frees g []) s of
+                NONE => sys_error "try_param_tac: no such variable"
+              | SOME T => [(P, if ca then concl else lambda (Free (s, T)) concl), (x, Free (s, T))])
+          | (_, T) :: _ =>
+              [(P, list_abs (params, if ca then concl else incr_boundvars 1 (Abs (s, T, concl)))),
+                (x, list_abs (params, Bound 0))])) rule';
   in compose_tac (false, rule'', nprems_of rule) i end);
 
 
@@ -1635,7 +1611,8 @@
         val suff = if i = 0 then ext_typeN else inner_typeN ^ string_of_int i;
         val ((_, cons), thy') = thy
           |> Iso_Tuple_Support.add_iso_tuple_type
-            (Binding.suffix_name suff (Binding.name base_name), alphas_zeta) (fastype_of left, fastype_of right);
+            (Binding.suffix_name suff (Binding.name base_name), alphas_zeta)
+              (fastype_of left, fastype_of right);
       in
         (cons $ left $ right, (thy', i + 1))
       end;
@@ -1778,7 +1755,10 @@
             ("ext_surjective", surject),
             ("ext_split", split_meta)]);
 
-  in (((ext_name, ext_type), (ext_tyco, alphas_zeta), extT, induct', inject', surjective', split_meta', ext_def), thm_thy) end;
+  in
+    (((ext_name, ext_type), (ext_tyco, alphas_zeta),
+      extT, induct', inject', surjective', split_meta', ext_def), thm_thy)
+  end;
 
 fun chunks [] [] = []
   | chunks [] xs = [xs]
@@ -1795,7 +1775,7 @@
 
 (* mk_recordT *)
 
-(*builds up the record type from the current extension tpye extT and a list
+(*build up the record type from the current extension tpye extT and a list
   of parent extensions, starting with the root of the record hierarchy*)
 fun mk_recordT extT =
   fold_rev (fn (parent, Ts) => fn T => Type (parent, subst_last T Ts)) extT;
@@ -1834,8 +1814,10 @@
     val lhs = HOLogic.mk_random T size;
     val tc = HOLogic.mk_return Tm @{typ Random.seed}
       (HOLogic.mk_valtermify_app extN params T);
-    val rhs = HOLogic.mk_ST
-      (map (fn (v, T') => ((HOLogic.mk_random T' size, @{typ Random.seed}), SOME (v, termifyT T'))) params)
+    val rhs =
+      HOLogic.mk_ST
+        (map (fn (v, T') =>
+          ((HOLogic.mk_random T' size, @{typ Random.seed}), SOME (v, termifyT T'))) params)
         tc @{typ Random.seed} (SOME Tm, @{typ Random.seed});
     val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
   in 
@@ -1860,23 +1842,26 @@
 
 fun add_code ext_tyco vs extT ext simps inject thy =
   let
-    val eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq)
-      (Const (@{const_name eq_class.eq}, extT --> extT --> HOLogic.boolT),
-       Const (@{const_name "op ="}, extT --> extT --> HOLogic.boolT));
-    fun tac eq_def = Class.intro_classes_tac []
+    val eq =
+      (HOLogic.mk_Trueprop o HOLogic.mk_eq)
+        (Const (@{const_name HOL.equal}, extT --> extT --> HOLogic.boolT),
+         Const (@{const_name HOL.eq}, extT --> extT --> HOLogic.boolT));
+    fun tac eq_def =
+      Class.intro_classes_tac []
       THEN (Simplifier.rewrite_goals_tac [Simpdata.mk_eq eq_def])
       THEN ALLGOALS (rtac @{thm refl});
     fun mk_eq thy eq_def = Simplifier.rewrite_rule
       [(AxClass.unoverload thy o Thm.symmetric o Simpdata.mk_eq) eq_def] inject;
-    fun mk_eq_refl thy = @{thm HOL.eq_refl}
+    fun mk_eq_refl thy =
+      @{thm equal_refl}
       |> Thm.instantiate
-         ([pairself (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort eq}), Logic.varifyT_global extT)], [])
+        ([pairself (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort equal}), Logic.varifyT_global extT)], [])
       |> AxClass.unoverload thy;
   in
     thy
     |> Code.add_datatype [ext]
     |> fold Code.add_default_eqn simps
-    |> Class.instantiation ([ext_tyco], vs, [HOLogic.class_eq])
+    |> Class.instantiation ([ext_tyco], vs, [HOLogic.class_equal])
     |> `(fn lthy => Syntax.check_term lthy eq)
     |-> (fn eq => Specification.definition
          (NONE, (Attrib.empty_binding, eq)))
@@ -1944,7 +1929,8 @@
 
     val extension_name = full binding;
 
-    val ((ext, (ext_tyco, vs), extT, ext_induct, ext_inject, ext_surjective, ext_split, ext_def), ext_thy) =
+    val ((ext, (ext_tyco, vs),
+        extT, ext_induct, ext_inject, ext_surjective, ext_split, ext_def), ext_thy) =
       thy
       |> Sign.qualified_path false binding
       |> extension_definition extension_name fields alphas_ext zeta moreT more vars;
--- a/src/HOL/Tools/refute.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/refute.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -647,10 +647,10 @@
       | Const (@{const_name Hilbert_Choice.Eps}, _) => t
       | Const (@{const_name All}, _) => t
       | Const (@{const_name Ex}, _) => t
-      | Const (@{const_name "op ="}, _) => t
-      | Const (@{const_name "op &"}, _) => t
-      | Const (@{const_name "op |"}, _) => t
-      | Const (@{const_name "op -->"}, _) => t
+      | Const (@{const_name HOL.eq}, _) => t
+      | Const (@{const_name HOL.conj}, _) => t
+      | Const (@{const_name HOL.disj}, _) => t
+      | Const (@{const_name HOL.implies}, _) => t
       (* sets *)
       | Const (@{const_name Collect}, _) => t
       | Const (@{const_name Set.member}, _) => t
@@ -823,10 +823,10 @@
         end
       | Const (@{const_name All}, T) => collect_type_axioms T axs
       | Const (@{const_name Ex}, T) => collect_type_axioms T axs
-      | Const (@{const_name "op ="}, T) => collect_type_axioms T axs
-      | Const (@{const_name "op &"}, _) => axs
-      | Const (@{const_name "op |"}, _) => axs
-      | Const (@{const_name "op -->"}, _) => axs
+      | Const (@{const_name HOL.eq}, T) => collect_type_axioms T axs
+      | Const (@{const_name HOL.conj}, _) => axs
+      | Const (@{const_name HOL.disj}, _) => axs
+      | Const (@{const_name HOL.implies}, _) => axs
       (* sets *)
       | Const (@{const_name Collect}, T) => collect_type_axioms T axs
       | Const (@{const_name Set.member}, T) => collect_type_axioms T axs
@@ -1850,18 +1850,18 @@
       end
     | Const (@{const_name Ex}, _) =>
       SOME (interpret thy model args (eta_expand t 1))
-    | Const (@{const_name "op ="}, _) $ t1 $ t2 =>  (* similar to "==" (Pure) *)
+    | Const (@{const_name HOL.eq}, _) $ t1 $ t2 =>  (* similar to "==" (Pure) *)
       let
         val (i1, m1, a1) = interpret thy model args t1
         val (i2, m2, a2) = interpret thy m1 a1 t2
       in
         SOME (make_equality (i1, i2), m2, a2)
       end
-    | Const (@{const_name "op ="}, _) $ t1 =>
+    | Const (@{const_name HOL.eq}, _) $ t1 =>
       SOME (interpret thy model args (eta_expand t 1))
-    | Const (@{const_name "op ="}, _) =>
+    | Const (@{const_name HOL.eq}, _) =>
       SOME (interpret thy model args (eta_expand t 2))
-    | Const (@{const_name "op &"}, _) $ t1 $ t2 =>
+    | Const (@{const_name HOL.conj}, _) $ t1 $ t2 =>
       (* 3-valued logic *)
       let
         val (i1, m1, a1) = interpret thy model args t1
@@ -1871,14 +1871,14 @@
       in
         SOME (Leaf [fmTrue, fmFalse], m2, a2)
       end
-    | Const (@{const_name "op &"}, _) $ t1 =>
+    | Const (@{const_name HOL.conj}, _) $ t1 =>
       SOME (interpret thy model args (eta_expand t 1))
-    | Const (@{const_name "op &"}, _) =>
+    | Const (@{const_name HOL.conj}, _) =>
       SOME (interpret thy model args (eta_expand t 2))
       (* this would make "undef" propagate, even for formulae like *)
       (* "False & undef":                                          *)
       (* SOME (Node [Node [TT, FF], Node [FF, FF]], model, args) *)
-    | Const (@{const_name "op |"}, _) $ t1 $ t2 =>
+    | Const (@{const_name HOL.disj}, _) $ t1 $ t2 =>
       (* 3-valued logic *)
       let
         val (i1, m1, a1) = interpret thy model args t1
@@ -1888,14 +1888,14 @@
       in
         SOME (Leaf [fmTrue, fmFalse], m2, a2)
       end
-    | Const (@{const_name "op |"}, _) $ t1 =>
+    | Const (@{const_name HOL.disj}, _) $ t1 =>
       SOME (interpret thy model args (eta_expand t 1))
-    | Const (@{const_name "op |"}, _) =>
+    | Const (@{const_name HOL.disj}, _) =>
       SOME (interpret thy model args (eta_expand t 2))
       (* this would make "undef" propagate, even for formulae like *)
       (* "True | undef":                                           *)
       (* SOME (Node [Node [TT, TT], Node [TT, FF]], model, args) *)
-    | Const (@{const_name "op -->"}, _) $ t1 $ t2 =>  (* similar to "==>" (Pure) *)
+    | Const (@{const_name HOL.implies}, _) $ t1 $ t2 =>  (* similar to "==>" (Pure) *)
       (* 3-valued logic *)
       let
         val (i1, m1, a1) = interpret thy model args t1
@@ -1905,9 +1905,9 @@
       in
         SOME (Leaf [fmTrue, fmFalse], m2, a2)
       end
-    | Const (@{const_name "op -->"}, _) $ t1 =>
+    | Const (@{const_name HOL.implies}, _) $ t1 =>
       SOME (interpret thy model args (eta_expand t 1))
-    | Const (@{const_name "op -->"}, _) =>
+    | Const (@{const_name HOL.implies}, _) =>
       SOME (interpret thy model args (eta_expand t 2))
       (* this would make "undef" propagate, even for formulae like *)
       (* "False --> undef":                                        *)
--- a/src/HOL/Tools/simpdata.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/simpdata.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -10,11 +10,11 @@
 structure Quantifier1 = Quantifier1Fun
 (struct
   (*abstract syntax*)
-  fun dest_eq ((c as Const(@{const_name "op ="},_)) $ s $ t) = SOME (c, s, t)
+  fun dest_eq ((c as Const(@{const_name HOL.eq},_)) $ s $ t) = SOME (c, s, t)
     | dest_eq _ = NONE;
-  fun dest_conj ((c as Const(@{const_name "op &"},_)) $ s $ t) = SOME (c, s, t)
+  fun dest_conj ((c as Const(@{const_name HOL.conj},_)) $ s $ t) = SOME (c, s, t)
     | dest_conj _ = NONE;
-  fun dest_imp ((c as Const(@{const_name "op -->"},_)) $ s $ t) = SOME (c, s, t)
+  fun dest_imp ((c as Const(@{const_name HOL.implies},_)) $ s $ t) = SOME (c, s, t)
     | dest_imp _ = NONE;
   val conj = HOLogic.conj
   val imp  = HOLogic.imp
@@ -44,7 +44,7 @@
 fun mk_eq th = case concl_of th
   (*expects Trueprop if not == *)
   of Const (@{const_name "=="},_) $ _ $ _ => th
-   | _ $ (Const (@{const_name "op ="}, _) $ _ $ _) => mk_meta_eq th
+   | _ $ (Const (@{const_name HOL.eq}, _) $ _ $ _) => mk_meta_eq th
    | _ $ (Const (@{const_name Not}, _) $ _) => th RS @{thm Eq_FalseI}
    | _ => th RS @{thm Eq_TrueI}
 
@@ -159,8 +159,8 @@
   (Scan.succeed "Clasimp.clasimpset_of (ML_Context.the_local_context ())");
 
 val mksimps_pairs =
- [(@{const_name "op -->"}, [@{thm mp}]),
-  (@{const_name "op &"}, [@{thm conjunct1}, @{thm conjunct2}]),
+ [(@{const_name HOL.implies}, [@{thm mp}]),
+  (@{const_name HOL.conj}, [@{thm conjunct1}, @{thm conjunct2}]),
   (@{const_name All}, [@{thm spec}]),
   (@{const_name True}, []),
   (@{const_name False}, []),
--- a/src/HOL/Tools/string_code.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/string_code.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -59,7 +59,7 @@
                   Code_Printer.literal_list literals (map (pr vars Code_Printer.NOBR) ts))
         | NONE =>
             List_Code.default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;
-  in Code_Target.add_syntax_const target
+  in Code_Target.add_const_syntax target
     @{const_name Cons} (SOME (Code_Printer.complex_const_syntax (2, (cs_summa, pretty))))
   end;
 
@@ -69,7 +69,7 @@
       case decode_char nibbles' (t1, t2)
        of SOME c => (Code_Printer.str o Code_Printer.literal_char literals) c
         | NONE => Code_Printer.eqn_error thm "Illegal character expression";
-  in Code_Target.add_syntax_const target
+  in Code_Target.add_const_syntax target
     @{const_name Char} (SOME (Code_Printer.complex_const_syntax (2, (cs_nibbles, pretty))))
   end;
 
@@ -82,7 +82,7 @@
              of SOME p => p
               | NONE => Code_Printer.eqn_error thm "Illegal message expression")
         | NONE => Code_Printer.eqn_error thm "Illegal message expression";
-  in Code_Target.add_syntax_const target 
+  in Code_Target.add_const_syntax target 
     @{const_name STR} (SOME (Code_Printer.complex_const_syntax (1, (cs_summa, pretty))))
   end;
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/try.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -0,0 +1,81 @@
+(*  Title:      HOL/Tools/try.ML
+    Author:     Jasmin Blanchette, TU Muenchen
+
+Try a combination of proof methods.
+*)
+
+signature TRY =
+sig
+  val invoke_try : Proof.state -> unit
+end;
+
+structure Try : TRY =
+struct
+
+val timeout = Time.fromSeconds 5
+
+fun can_apply pre post tac st =
+  let val {goal, ...} = Proof.goal st in
+    case TimeLimit.timeLimit timeout (Seq.pull o tac) (pre st) of
+      SOME (x, _) => nprems_of (post x) < nprems_of goal
+    | NONE => false
+  end
+
+fun do_generic command pre post apply st =
+  let val timer = Timer.startRealTimer () in
+    if can_apply pre post apply st then
+      SOME (command, Time.toMilliseconds (Timer.checkRealTimer timer))
+    else
+      NONE
+  end
+
+fun named_method thy name =
+  Method.method thy (Args.src ((name, []), Position.none))
+
+fun apply_named_method name ctxt =
+  let val thy = ProofContext.theory_of ctxt in
+    Method.apply (named_method thy name) ctxt []
+  end
+
+fun do_named_method name st =
+  do_generic name (#goal o Proof.goal) snd
+             (apply_named_method name (Proof.context_of st)) st
+
+fun apply_named_method_on_first_goal name ctxt =
+  let val thy = ProofContext.theory_of ctxt in
+    Proof.refine (Method.SelectGoals (1, Method.Basic (named_method thy name)))
+  end
+
+fun do_named_method_on_first_goal name st =
+  do_generic (name ^ (if nprems_of (#goal (Proof.goal st)) > 1 then "[1]"
+                      else "")) I (#goal o Proof.goal)
+             (apply_named_method_on_first_goal name (Proof.context_of st)) st
+
+val all_goals_named_methods = ["auto", "safe"]
+val first_goal_named_methods =
+  ["simp", "fast", "fastsimp", "force", "best", "blast"]
+val do_methods =
+  map do_named_method_on_first_goal all_goals_named_methods @
+  map do_named_method first_goal_named_methods
+
+fun time_string (s, ms) = s ^ ": " ^ string_of_int ms ^ " ms"
+
+fun invoke_try st =
+  case do_methods |> Par_List.map (fn f => f st)
+                  |> map_filter I |> sort (int_ord o pairself snd) of
+    [] => writeln "No proof found."
+  | xs as (s, _) :: _ =>
+    priority ("Try this command: " ^
+        Markup.markup Markup.sendback
+            ((if nprems_of (#goal (Proof.goal st)) = 1 then "by" else "apply") ^
+             " " ^ s) ^
+        ".\n(" ^ space_implode "; " (map time_string xs) ^ ")\n")
+
+val tryN = "try"
+
+val _ =
+  Outer_Syntax.improper_command tryN
+      "try a combination of proof methods" Keyword.diag
+      (Scan.succeed (Toplevel.keep (invoke_try o Toplevel.proof_of)))
+
+end;
--- a/src/HOL/Tools/typedef.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/typedef.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -186,7 +186,8 @@
           ||> Thm.term_of;
 
         val ((RepC, AbsC, axiom_name, axiom), axiom_lthy) = set_lthy |>
-          Local_Theory.theory_result (primitive_typedef typedef_name newT oldT Rep_name Abs_name A);
+          Local_Theory.background_theory_result
+            (primitive_typedef typedef_name newT oldT Rep_name Abs_name A);
 
         val cert = Thm.cterm_of (ProofContext.theory_of axiom_lthy);
         val typedef =
@@ -246,7 +247,7 @@
       in
         lthy2
         |> Local_Theory.declaration true (fn phi => put_info full_tname (transform_info phi info))
-        |> Local_Theory.theory (Typedef_Interpretation.data full_tname)
+        |> Local_Theory.background_theory (Typedef_Interpretation.data full_tname)
         |> pair (full_tname, info)
       end;
 
--- a/src/HOL/Typerep.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Typerep.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -78,9 +78,13 @@
 *}
 
 lemma [code]:
-  "eq_class.eq (Typerep tyco1 tys1) (Typerep tyco2 tys2) \<longleftrightarrow> eq_class.eq tyco1 tyco2
-     \<and> list_all2 eq_class.eq tys1 tys2"
-  by (auto simp add: equals_eq [symmetric] list_all2_eq [symmetric])
+  "HOL.equal (Typerep tyco1 tys1) (Typerep tyco2 tys2) \<longleftrightarrow> HOL.equal tyco1 tyco2
+     \<and> list_all2 HOL.equal tys1 tys2"
+  by (auto simp add: eq_equal [symmetric] list_all2_eq [symmetric])
+
+lemma [code nbe]:
+  "HOL.equal (x :: typerep) x \<longleftrightarrow> True"
+  by (fact equal_refl)
 
 code_type typerep
   (Eval "Term.typ")
--- a/src/HOL/Word/Word.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Word/Word.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -953,14 +953,14 @@
 
 text {* Executable equality *}
 
-instantiation word :: ("{len0}") eq
+instantiation word :: (len0) equal
 begin
 
-definition eq_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool" where
-  "eq_word k l \<longleftrightarrow> HOL.eq (uint k) (uint l)"
+definition equal_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool" where
+  "equal_word k l \<longleftrightarrow> HOL.equal (uint k) (uint l)"
 
 instance proof
-qed (simp add: eq eq_word_def)
+qed (simp add: equal equal_word_def)
 
 end
 
--- a/src/HOL/ex/Meson_Test.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/ex/Meson_Test.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -10,7 +10,7 @@
   below and constants declared in HOL!
 *}
 
-hide_const (open) subset quotient union inter sum
+hide_const (open) implies union inter subset sum quotient 
 
 text {*
   Test data for the MESON proof procedure
--- a/src/HOL/ex/Numeral.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/ex/Numeral.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -97,7 +97,7 @@
 structure Dig_Simps = Named_Thms
 (
   val name = "numeral"
-  val description = "Simplification rules for numerals"
+  val description = "simplification rules for numerals"
 )
 *}
 
@@ -845,7 +845,7 @@
   "(uminus :: int \<Rightarrow> int) = uminus"
   "(op - :: int \<Rightarrow> int \<Rightarrow> int) = op -"
   "(op * :: int \<Rightarrow> int \<Rightarrow> int) = op *"
-  "(eq_class.eq :: int \<Rightarrow> int \<Rightarrow> bool) = eq_class.eq"
+  "(HOL.equal :: int \<Rightarrow> int \<Rightarrow> bool) = HOL.equal"
   "(op \<le> :: int \<Rightarrow> int \<Rightarrow> bool) = op \<le>"
   "(op < :: int \<Rightarrow> int \<Rightarrow> bool) = op <"
   by rule+
@@ -928,16 +928,20 @@
   by simp_all
 
 lemma eq_int_code [code]:
-  "eq_class.eq 0 (0::int) \<longleftrightarrow> True"
-  "eq_class.eq 0 (Pls l) \<longleftrightarrow> False"
-  "eq_class.eq 0 (Mns l) \<longleftrightarrow> False"
-  "eq_class.eq (Pls k) 0 \<longleftrightarrow> False"
-  "eq_class.eq (Pls k) (Pls l) \<longleftrightarrow> eq_class.eq k l"
-  "eq_class.eq (Pls k) (Mns l) \<longleftrightarrow> False"
-  "eq_class.eq (Mns k) 0 \<longleftrightarrow> False"
-  "eq_class.eq (Mns k) (Pls l) \<longleftrightarrow> False"
-  "eq_class.eq (Mns k) (Mns l) \<longleftrightarrow> eq_class.eq k l"
-  by (auto simp add: eq dest: sym)
+  "HOL.equal 0 (0::int) \<longleftrightarrow> True"
+  "HOL.equal 0 (Pls l) \<longleftrightarrow> False"
+  "HOL.equal 0 (Mns l) \<longleftrightarrow> False"
+  "HOL.equal (Pls k) 0 \<longleftrightarrow> False"
+  "HOL.equal (Pls k) (Pls l) \<longleftrightarrow> HOL.equal k l"
+  "HOL.equal (Pls k) (Mns l) \<longleftrightarrow> False"
+  "HOL.equal (Mns k) 0 \<longleftrightarrow> False"
+  "HOL.equal (Mns k) (Pls l) \<longleftrightarrow> False"
+  "HOL.equal (Mns k) (Mns l) \<longleftrightarrow> HOL.equal k l"
+  by (auto simp add: equal dest: sym)
+
+lemma [code nbe]:
+  "HOL.equal (k::int) k \<longleftrightarrow> True"
+  by (fact equal_refl)
 
 lemma less_eq_int_code [code]:
   "0 \<le> (0::int) \<longleftrightarrow> True"
@@ -985,7 +989,7 @@
       in dest_num end;
     fun pretty sgn literals [one', dig0', dig1'] _ thm _ _ [(t, _)] =
       (Code_Printer.str o print literals o sgn o dest_num one' dig0' dig1' thm) t
-    fun add_syntax (c, sgn) = Code_Target.add_syntax_const target c
+    fun add_syntax (c, sgn) = Code_Target.add_const_syntax target c
       (SOME (Code_Printer.complex_const_syntax
         (1, ([@{const_name One}, @{const_name Dig0}, @{const_name Dig1}],
           pretty sgn))));
@@ -1033,14 +1037,14 @@
   (SML "IntInf.- ((_), 1)")
   (OCaml "Big'_int.pred'_big'_int")
   (Haskell "!(_/ -/ 1)")
-  (Scala "!(_/ -/ 1)")
+  (Scala "!(_ -/ 1)")
   (Eval "!(_/ -/ 1)")
 
 code_const Int.succ
   (SML "IntInf.+ ((_), 1)")
   (OCaml "Big'_int.succ'_big'_int")
   (Haskell "!(_/ +/ 1)")
-  (Scala "!(_/ +/ 1)")
+  (Scala "!(_ +/ 1)")
   (Eval "!(_/ +/ 1)")
 
 code_const "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int"
@@ -1078,7 +1082,7 @@
   (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))")
   (Eval "Integer.div'_mod/ (abs _)/ (abs _)")
 
-code_const "eq_class.eq \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
   (SML "!((_ : IntInf.int) = _)")
   (OCaml "Big'_int.eq'_big'_int")
   (Haskell infixl 4 "==")
--- a/src/HOL/ex/SAT_Examples.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/ex/SAT_Examples.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -10,8 +10,8 @@
 
 (* ML {* sat.solver := "zchaff_with_proofs"; *} *)
 (* ML {* sat.solver := "minisat_with_proofs"; *} *)
-ML {* Unsynchronized.set sat.trace_sat; *}
-ML {* Unsynchronized.set quick_and_dirty; *}
+ML {* sat.trace_sat := true; *}
+ML {* quick_and_dirty := true; *}
 
 lemma "True"
 by sat
@@ -77,8 +77,8 @@
 lemma "(ALL x. P x) | ~ All P"
 by sat
 
-ML {* Unsynchronized.reset sat.trace_sat; *}
-ML {* Unsynchronized.reset quick_and_dirty; *}
+ML {* sat.trace_sat := false; *}
+ML {* quick_and_dirty := false; *}
 
 method_setup rawsat =
   {* Scan.succeed (SIMPLE_METHOD' o sat.rawsat_tac) *}
@@ -525,8 +525,8 @@
 
 (* ML {*
 sat.solver := "zchaff_with_proofs";
-Unsynchronized.reset sat.trace_sat;
-Unsynchronized.reset quick_and_dirty;
+sat.trace_sat := false;
+quick_and_dirty := false;
 *} *)
 
 ML {*
--- a/src/HOL/ex/SVC_Oracle.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/ex/SVC_Oracle.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -88,13 +88,13 @@
             else if T = HOLogic.boolT then c $ (fm x) $ (fm y)
             else replace (c $ x $ y)   (*non-numeric comparison*)
     (*abstraction of a formula*)
-    and fm ((c as Const(@{const_name "op &"}, _)) $ p $ q) = c $ (fm p) $ (fm q)
-      | fm ((c as Const(@{const_name "op |"}, _)) $ p $ q) = c $ (fm p) $ (fm q)
-      | fm ((c as Const(@{const_name "op -->"}, _)) $ p $ q) = c $ (fm p) $ (fm q)
+    and fm ((c as Const(@{const_name HOL.conj}, _)) $ p $ q) = c $ (fm p) $ (fm q)
+      | fm ((c as Const(@{const_name HOL.disj}, _)) $ p $ q) = c $ (fm p) $ (fm q)
+      | fm ((c as Const(@{const_name HOL.implies}, _)) $ p $ q) = c $ (fm p) $ (fm q)
       | fm ((c as Const(@{const_name Not}, _)) $ p) = c $ (fm p)
       | fm ((c as Const(@{const_name True}, _))) = c
       | fm ((c as Const(@{const_name False}, _))) = c
-      | fm (t as Const(@{const_name "op ="},  Type ("fun", [T,_])) $ _ $ _) = rel (T, t)
+      | fm (t as Const(@{const_name HOL.eq},  Type ("fun", [T,_])) $ _ $ _) = rel (T, t)
       | fm (t as Const(@{const_name Orderings.less},  Type ("fun", [T,_])) $ _ $ _) = rel (T, t)
       | fm (t as Const(@{const_name Orderings.less_eq}, Type ("fun", [T,_])) $ _ $ _) = rel (T, t)
       | fm t = replace t
--- a/src/HOL/ex/svc_funcs.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/ex/svc_funcs.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -89,13 +89,13 @@
    let fun tag t =
          let val (c,ts) = strip_comb t
          in  case c of
-             Const(@{const_name "op &"}, _)   => apply c (map tag ts)
-           | Const(@{const_name "op |"}, _)   => apply c (map tag ts)
-           | Const(@{const_name "op -->"}, _) => apply c (map tag ts)
+             Const(@{const_name HOL.conj}, _)   => apply c (map tag ts)
+           | Const(@{const_name HOL.disj}, _)   => apply c (map tag ts)
+           | Const(@{const_name HOL.implies}, _) => apply c (map tag ts)
            | Const(@{const_name Not}, _)    => apply c (map tag ts)
            | Const(@{const_name True}, _)   => (c, false)
            | Const(@{const_name False}, _)  => (c, false)
-           | Const(@{const_name "op ="}, Type ("fun", [T,_])) =>
+           | Const(@{const_name HOL.eq}, Type ("fun", [T,_])) =>
                  if T = HOLogic.boolT then
                      (*biconditional: with int/nat comparisons below?*)
                      let val [t1,t2] = ts
@@ -183,11 +183,11 @@
       | tm t = Int (lit t)
                handle Match => var (t,[])
     (*translation of a formula*)
-    and fm pos (Const(@{const_name "op &"}, _) $ p $ q) =
+    and fm pos (Const(@{const_name HOL.conj}, _) $ p $ q) =
             Buildin("AND", [fm pos p, fm pos q])
-      | fm pos (Const(@{const_name "op |"}, _) $ p $ q) =
+      | fm pos (Const(@{const_name HOL.disj}, _) $ p $ q) =
             Buildin("OR", [fm pos p, fm pos q])
-      | fm pos (Const(@{const_name "op -->"}, _) $ p $ q) =
+      | fm pos (Const(@{const_name HOL.implies}, _) $ p $ q) =
             Buildin("=>", [fm (not pos) p, fm pos q])
       | fm pos (Const(@{const_name Not}, _) $ p) =
             Buildin("NOT", [fm (not pos) p])
@@ -200,7 +200,7 @@
             Buildin("AND",   (*unfolding uses both polarities*)
                          [Buildin("=>", [fm (not pos) p, fm pos q]),
                           Buildin("=>", [fm (not pos) q, fm pos p])])
-      | fm pos (t as Const(@{const_name "op ="}, Type ("fun", [T,_])) $ x $ y) =
+      | fm pos (t as Const(@{const_name HOL.eq}, Type ("fun", [T,_])) $ x $ y) =
             let val tx = tm x and ty = tm y
                 in if pos orelse T = HOLogic.realT then
                        Buildin("=", [tx, ty])
--- a/src/HOLCF/Tools/Domain/domain_library.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOLCF/Tools/Domain/domain_library.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -121,7 +121,7 @@
       val r_inst = read_instantiate ctxt;
       fun at thm =
           case concl_of thm of
-            _$(Const(@{const_name "op &"},_)$_$_)       => at(thm RS conjunct1)@at(thm RS conjunct2)
+            _$(Const(@{const_name HOL.conj},_)$_$_)       => at(thm RS conjunct1)@at(thm RS conjunct2)
           | _$(Const(@{const_name All} ,_)$Abs(s,_,_))=> at(thm RS (r_inst [(("x", 0), "?" ^ s)] spec))
           | _                             => [thm];
     in map zero_var_indexes (at thm) end;
@@ -190,9 +190,9 @@
 
 fun mk_All  (x,P) = %%:"all" $ mk_lam(x,P); (* meta universal quantification *)
 
-infixr 0 ===>;  fun S ===> T = %%:"==>" $ S $ T;
-infix 0 ==;     fun S ==  T = %%:"==" $ S $ T;
-infix 1 ===;    fun S === T = %%:"op =" $ S $ T;
+infixr 0 ===>;  fun S ===> T = %%: "==>" $ S $ T;
+infix 0 ==;     fun S ==  T = %%: "==" $ S $ T;
+infix 1 ===;    fun S === T = %%: @{const_name HOL.eq} $ S $ T;
 infix 1 ~=;     fun S ~=  T = HOLogic.mk_not (S === T);
 
 infix 9 `  ; fun f ` x = %%: @{const_name Rep_CFun} $ f $ x;
--- a/src/HOLCF/Tools/pcpodef.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOLCF/Tools/pcpodef.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -326,7 +326,7 @@
     val args = map (apsnd (prep_constraint ctxt)) raw_args;
     val (goal1, goal2, make_result) =
       prepare_cpodef prep_term def name (b, args, mx) set opt_morphs thy;
-    fun after_qed [[th1, th2]] = ProofContext.theory (snd o make_result th1 th2)
+    fun after_qed [[th1, th2]] = ProofContext.background_theory (snd o make_result th1 th2)
       | after_qed _ = raise Fail "cpodef_proof";
   in Proof.theorem NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end;
 
@@ -337,7 +337,7 @@
     val args = map (apsnd (prep_constraint ctxt)) raw_args;
     val (goal1, goal2, make_result) =
       prepare_pcpodef prep_term def name (b, args, mx) set opt_morphs thy;
-    fun after_qed [[th1, th2]] = ProofContext.theory (snd o make_result th1 th2)
+    fun after_qed [[th1, th2]] = ProofContext.background_theory (snd o make_result th1 th2)
       | after_qed _ = raise Fail "pcpodef_proof";
   in Proof.theorem NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end;
 
--- a/src/Provers/Arith/fast_lin_arith.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Provers/Arith/fast_lin_arith.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -88,13 +88,19 @@
   val cut_lin_arith_tac: simpset -> int -> tactic
   val lin_arith_tac: Proof.context -> bool -> int -> tactic
   val lin_arith_simproc: simpset -> term -> thm option
-  val map_data: ({add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list,
-                 lessD: thm list, neqE: thm list, simpset: Simplifier.simpset,
-                 number_of : serial * (theory -> typ -> int -> cterm)}
-                 -> {add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list,
-                     lessD: thm list, neqE: thm list, simpset: Simplifier.simpset,
-                     number_of : serial * (theory -> typ -> int -> cterm)})
-                -> Context.generic -> Context.generic
+  val map_data:
+    ({add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list,
+      lessD: thm list, neqE: thm list, simpset: Simplifier.simpset,
+      number_of: (theory -> typ -> int -> cterm) option} ->
+     {add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list,
+      lessD: thm list, neqE: thm list, simpset: Simplifier.simpset,
+      number_of: (theory -> typ -> int -> cterm) option}) ->
+      Context.generic -> Context.generic
+  val add_inj_thms: thm list -> Context.generic -> Context.generic
+  val add_lessD: thm -> Context.generic -> Context.generic
+  val add_simps: thm list -> Context.generic -> Context.generic
+  val add_simprocs: simproc list -> Context.generic -> Context.generic
+  val set_number_of: (theory -> typ -> int -> cterm) -> Context.generic -> Context.generic
   val trace: bool Unsynchronized.ref
 end;
 
@@ -105,8 +111,6 @@
 
 (** theory data **)
 
-fun no_number_of _ _ _ = raise CTERM ("number_of", [])
-
 structure Data = Generic_Data
 (
   type T =
@@ -116,32 +120,57 @@
     lessD: thm list,
     neqE: thm list,
     simpset: Simplifier.simpset,
-    number_of : serial * (theory -> typ -> int -> cterm)};
+    number_of: (theory -> typ -> int -> cterm) option};
 
-  val empty = {add_mono_thms = [], mult_mono_thms = [], inj_thms = [],
-               lessD = [], neqE = [], simpset = Simplifier.empty_ss,
-               number_of = (serial (), no_number_of) } : T;
+  val empty : T =
+   {add_mono_thms = [], mult_mono_thms = [], inj_thms = [],
+    lessD = [], neqE = [], simpset = Simplifier.empty_ss,
+    number_of = NONE};
   val extend = I;
   fun merge
-    ({add_mono_thms= add_mono_thms1, mult_mono_thms= mult_mono_thms1, inj_thms= inj_thms1,
-      lessD = lessD1, neqE=neqE1, simpset = simpset1,
-      number_of = (number_of1 as (s1, _))},
-     {add_mono_thms= add_mono_thms2, mult_mono_thms= mult_mono_thms2, inj_thms= inj_thms2,
-      lessD = lessD2, neqE=neqE2, simpset = simpset2,
-      number_of = (number_of2 as (s2, _))}) =
+    ({add_mono_thms = add_mono_thms1, mult_mono_thms = mult_mono_thms1, inj_thms = inj_thms1,
+      lessD = lessD1, neqE = neqE1, simpset = simpset1, number_of = number_of1},
+     {add_mono_thms = add_mono_thms2, mult_mono_thms = mult_mono_thms2, inj_thms = inj_thms2,
+      lessD = lessD2, neqE = neqE2, simpset = simpset2, number_of = number_of2}) : T =
     {add_mono_thms = Thm.merge_thms (add_mono_thms1, add_mono_thms2),
      mult_mono_thms = Thm.merge_thms (mult_mono_thms1, mult_mono_thms2),
      inj_thms = Thm.merge_thms (inj_thms1, inj_thms2),
      lessD = Thm.merge_thms (lessD1, lessD2),
      neqE = Thm.merge_thms (neqE1, neqE2),
      simpset = Simplifier.merge_ss (simpset1, simpset2),
-     (* FIXME depends on accidental load order !?! *)
-     number_of = if s1 > s2 then number_of1 else number_of2};
+     number_of = if is_some number_of1 then number_of1 else number_of2};
 );
 
 val map_data = Data.map;
 val get_data = Data.get o Context.Proof;
 
+fun map_inj_thms f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
+  {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = f inj_thms,
+    lessD = lessD, neqE = neqE, simpset = simpset, number_of = number_of};
+
+fun map_lessD f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
+  {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms,
+    lessD = f lessD, neqE = neqE, simpset = simpset, number_of = number_of};
+
+fun map_simpset f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
+  {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms,
+    lessD = lessD, neqE = neqE, simpset = f simpset, number_of = number_of};
+
+fun add_inj_thms thms = map_data (map_inj_thms (append thms));
+fun add_lessD thm = map_data (map_lessD (fn thms => thms @ [thm]));
+fun add_simps thms = map_data (map_simpset (fn simpset => simpset addsimps thms));
+fun add_simprocs procs = map_data (map_simpset (fn simpset => simpset addsimprocs procs));
+
+fun set_number_of f =
+  map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, ...} =>
+   {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms,
+    lessD = lessD, neqE = neqE, simpset = simpset, number_of = SOME f});
+
+fun number_of ctxt =
+  (case Data.get (Context.Proof ctxt) of
+    {number_of = SOME f, ...} => f (ProofContext.theory_of ctxt)
+  | _ => fn _ => fn _ => raise CTERM ("number_of", []));
+
 
 
 (*** A fast decision procedure ***)
@@ -446,8 +475,8 @@
   let
     val ctxt = Simplifier.the_context ss;
     val thy = ProofContext.theory_of ctxt;
-    val {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset,
-      number_of = (_, num_of), ...} = get_data ctxt;
+    val {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset, ...} = get_data ctxt;
+    val number_of = number_of ctxt;
     val simpset' = Simplifier.inherit_context ss simpset;
     fun only_concl f thm =
       if Thm.no_prems thm then f (Thm.concl_of thm) else NONE;
@@ -493,7 +522,7 @@
             val T = #T (Thm.rep_cterm cv)
           in
             mth
-            |> Thm.instantiate ([], [(cv, num_of thy T n)])
+            |> Thm.instantiate ([], [(cv, number_of T n)])
             |> rewrite_concl
             |> discharge_prem
             handle CTERM _ => mult_by_add n thm
--- a/src/Pure/Concurrent/future.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Concurrent/future.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -28,6 +28,7 @@
 {
   def peek: Option[Exn.Result[A]]
   def is_finished: Boolean = peek.isDefined
+  def get_finished: A = { require(is_finished); Exn.release(peek.get) }
   def join: A
   def map[B](f: A => B): Future[B] = Future.fork { f(join) }
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Concurrent/volatile.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -0,0 +1,22 @@
+/*  Title:      Pure/Concurrent/volatile.scala
+    Author:     Makarius
+
+Volatile variables.
+*/
+
+package isabelle
+
+
+class Volatile[A](init: A)
+{
+  @volatile private var state: A = init
+  def peek(): A = state
+  def change(f: A => A) { state = f(state) }
+  def change_yield[B](f: A => (B, A)): B =
+  {
+    val (result, new_state) = f(state)
+    state = new_state
+    result
+  }
+}
+
--- a/src/Pure/General/markup.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/General/markup.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -14,8 +14,8 @@
   val properties: Properties.T -> T -> T
   val nameN: string
   val name: string -> T -> T
+  val kindN: string
   val bindingN: string val binding: string -> T
-  val kindN: string
   val entityN: string val entity: string -> T
   val defN: string
   val refN: string
@@ -99,6 +99,8 @@
   val command_spanN: string val command_span: string -> T
   val ignored_spanN: string val ignored_span: T
   val malformed_spanN: string val malformed_span: T
+  val subgoalsN: string
+  val proof_stateN: string val proof_state: int -> T
   val stateN: string val state: T
   val subgoalN: string val subgoal: T
   val sendbackN: string val sendback: T
@@ -113,8 +115,10 @@
   val assignN: string val assign: int -> T
   val editN: string val edit: int -> int -> T
   val pidN: string
+  val serialN: string
   val promptN: string val prompt: T
   val readyN: string val ready: T
+  val reportN: string val report: T
   val no_output: output * output
   val default_output: T -> output * output
   val add_mode: string -> (T -> output * output) -> unit
@@ -156,21 +160,17 @@
 fun markup_int elem prop = (elem, fn i => (elem, [(prop, print_int i)]): T);
 
 
-(* name *)
+(* misc properties *)
 
 val nameN = "name";
 fun name a = properties [(nameN, a)];
 
-val (bindingN, binding) = markup_string "binding" nameN;
-
-
-(* kind *)
-
 val kindN = "kind";
 
 
 (* formal entities *)
 
+val (bindingN, binding) = markup_string "binding" nameN;
 val (entityN, entity) = markup_string "entity" nameN;
 
 val defN = "def";
@@ -305,6 +305,9 @@
 
 (* toplevel *)
 
+val subgoalsN = "subgoals";
+val (proof_stateN, proof_state) = markup_int "proof_state" subgoalsN;
+
 val (stateN, state) = markup_elem "state";
 val (subgoalN, subgoal) = markup_elem "subgoal";
 val (sendbackN, sendback) = markup_elem "sendback";
@@ -335,10 +338,13 @@
 (* messages *)
 
 val pidN = "pid";
+val serialN = "serial";
 
 val (promptN, prompt) = markup_elem "prompt";
 val (readyN, ready) = markup_elem "ready";
 
+val (reportN, report) = markup_elem "report";
+
 
 
 (** print mode operations **)
--- a/src/Pure/General/markup.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/General/markup.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -9,7 +9,7 @@
 
 object Markup
 {
-  /* integers */
+  /* plain values */
 
   object Int {
     def apply(i: scala.Int): String = i.toString
@@ -26,25 +26,33 @@
   }
 
 
-  /* property values */
-
-  def get_string(name: String, props: List[(String, String)]): Option[String] =
-    props.find(p => p._1 == name).map(_._2)
+  /* named properties */
 
-  def get_long(name: String, props: List[(String, String)]): Option[scala.Long] =
+  class Property(val name: String)
   {
-    get_string(name, props) match {
-      case None => None
-      case Some(Long(i)) => Some(i)
-    }
+    def apply(value: String): List[(String, String)] = List((name, value))
+    def unapply(props: List[(String, String)]): Option[String] =
+      props.find(_._1 == name).map(_._2)
   }
 
-  def get_int(name: String, props: List[(String, String)]): Option[scala.Int] =
+  class Int_Property(name: String)
   {
-    get_string(name, props) match {
-      case None => None
-      case Some(Int(i)) => Some(i)
-    }
+    def apply(value: scala.Int): List[(String, String)] = List((name, Int(value)))
+    def unapply(props: List[(String, String)]): Option[Int] =
+      props.find(_._1 == name) match {
+        case None => None
+        case Some((_, value)) => Int.unapply(value)
+      }
+  }
+
+  class Long_Property(name: String)
+  {
+    def apply(value: scala.Long): List[(String, String)] = List((name, Long(value)))
+    def unapply(props: List[(String, String)]): Option[Long] =
+      props.find(_._1 == name) match {
+        case None => None
+        case Some((_, value)) => Long.unapply(value)
+      }
   }
 
 
@@ -53,7 +61,7 @@
   val Empty = Markup("", Nil)
 
 
-  /* name */
+  /* misc properties */
 
   val NAME = "name"
   val KIND = "kind"
@@ -61,6 +69,7 @@
 
   /* formal entities */
 
+  val BINDING = "binding"
   val ENTITY = "entity"
   val DEF = "def"
   val REF = "ref"
@@ -86,9 +95,9 @@
 
   /* pretty printing */
 
-  val INDENT = "indent"
+  val Indent = new Int_Property("indent")
   val BLOCK = "block"
-  val WIDTH = "width"
+  val Width = new Int_Property("width")
   val BREAK = "break"
 
 
@@ -188,6 +197,9 @@
 
   /* toplevel */
 
+  val SUBGOALS = "subgoals"
+  val PROOF_STATE = "proof_state"
+
   val STATE = "state"
   val SUBGOAL = "subgoal"
   val SENDBACK = "sendback"
@@ -215,6 +227,7 @@
   /* messages */
 
   val PID = "pid"
+  val Serial = new Long_Property("serial")
 
   val MESSAGE = "message"
   val CLASS = "class"
--- a/src/Pure/General/output.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/General/output.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -45,7 +45,6 @@
   val status: string -> unit
   val report: string -> unit
   val debugging: bool Unsynchronized.ref
-  val no_warnings_CRITICAL: ('a -> 'b) -> 'a -> 'b
   val debug: (unit -> string) -> unit
 end;
 
@@ -80,11 +79,8 @@
 
 (* output primitives -- normally NOT used directly!*)
 
-fun std_output s = NAMED_CRITICAL "IO" (fn () =>
-  (TextIO.output (TextIO.stdOut, s); TextIO.flushOut TextIO.stdOut));
-
-fun std_error s = NAMED_CRITICAL "IO" (fn () =>
-  (TextIO.output (TextIO.stdErr, s); TextIO.flushOut TextIO.stdErr));
+fun std_output s = (TextIO.output (TextIO.stdOut, s); TextIO.flushOut TextIO.stdOut);
+fun std_error s = (TextIO.output (TextIO.stdErr, s); TextIO.flushOut TextIO.stdErr);
 
 fun writeln_default "" = ()
   | writeln_default s = std_output (suffix "\n" s);
@@ -113,8 +109,6 @@
 
 fun legacy_feature s = warning ("Legacy feature! " ^ s);
 
-fun no_warnings_CRITICAL f = setmp_CRITICAL warning_fn (K ()) f;
-
 val debugging = Unsynchronized.ref false;
 fun debug s = if ! debugging then ! debug_fn (output (s ())) else ()
 
--- a/src/Pure/General/position.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/General/position.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -27,6 +27,7 @@
   val of_properties: Properties.T -> T
   val properties_of: T -> Properties.T
   val default_properties: T -> Properties.T -> Properties.T
+  val report_markup: T -> Markup.T
   val report_text: Markup.T -> T -> string -> unit
   val report: Markup.T -> T -> unit
   val str_of: T -> string
@@ -125,6 +126,8 @@
   if exists (member (op =) Markup.position_properties o #1) props then props
   else properties_of default @ props;
 
+fun report_markup pos = Markup.properties (properties_of pos) Markup.report;
+
 fun report_text markup (pos as Pos (count, _)) txt =
   if invalid_count count then ()
   else Output.report (Markup.markup (Markup.properties (properties_of pos) markup) txt);
--- a/src/Pure/General/position.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/General/position.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -11,22 +11,32 @@
 {
   type T = List[(String, String)]
 
-  def get_line(pos: T): Option[Int] = Markup.get_int(Markup.LINE, pos)
-  def get_column(pos: T): Option[Int] = Markup.get_int(Markup.COLUMN, pos)
-  def get_offset(pos: T): Option[Int] = Markup.get_int(Markup.OFFSET, pos)
-  def get_end_line(pos: T): Option[Int] = Markup.get_int(Markup.END_LINE, pos)
-  def get_end_column(pos: T): Option[Int] = Markup.get_int(Markup.END_COLUMN, pos)
-  def get_end_offset(pos: T): Option[Int] = Markup.get_int(Markup.END_OFFSET, pos)
-  def get_file(pos: T): Option[String] = Markup.get_string(Markup.FILE, pos)
-  def get_id(pos: T): Option[Long] = Markup.get_long(Markup.ID, pos)
+  val Line = new Markup.Int_Property(Markup.LINE)
+  val End_Line = new Markup.Int_Property(Markup.END_LINE)
+  val Offset = new Markup.Int_Property(Markup.OFFSET)
+  val End_Offset = new Markup.Int_Property(Markup.END_OFFSET)
+  val File = new Markup.Property(Markup.FILE)
+  val Id = new Markup.Long_Property(Markup.ID)
 
-  def get_range(pos: T): Option[Text.Range] =
-    (get_offset(pos), get_end_offset(pos)) match {
-      case (Some(start), Some(stop)) if start <= stop => Some(Text.Range(start, stop))
-      case (Some(start), None) => Some(Text.Range(start))
-      case (_, _) => None
-    }
+  object Range
+  {
+    def apply(range: Text.Range): T = Offset(range.start) ++ Offset(range.stop)
+    def unapply(pos: T): Option[Text.Range] =
+      (Offset.unapply(pos), End_Offset.unapply(pos)) match {
+        case (Some(start), Some(stop)) if start <= stop => Some(Text.Range(start, stop))
+        case (Some(start), None) => Some(Text.Range(start))
+        case _ => None
+      }
+  }
 
-  object Id { def unapply(pos: T): Option[Long] = get_id(pos) }
-  object Range { def unapply(pos: T): Option[Text.Range] = get_range(pos) }
+  object Id_Range
+  {
+    def unapply(pos: T): Option[(Long, Text.Range)] =
+      (pos, pos) match {
+        case (Id(id), Range(range)) => Some((id, range))
+        case _ => None
+      }
+  }
+
+  def purge(props: T): T = props.filterNot(p => Markup.POSITION_PROPERTIES(p._1))
 }
--- a/src/Pure/General/pretty.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/General/pretty.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -19,12 +19,11 @@
   object Block
   {
     def apply(i: Int, body: XML.Body): XML.Tree =
-      XML.Elem(Markup(Markup.BLOCK, List((Markup.INDENT, Markup.Int(i)))), body)
+      XML.Elem(Markup(Markup.BLOCK, Markup.Indent(i)), body)
 
     def unapply(tree: XML.Tree): Option[(Int, XML.Body)] =
       tree match {
-        case XML.Elem(
-          Markup(Markup.BLOCK, List((Markup.INDENT, Markup.Int(i)))), body) => Some((i, body))
+        case XML.Elem(Markup(Markup.BLOCK, Markup.Indent(i)), body) => Some((i, body))
         case _ => None
       }
   }
@@ -32,12 +31,11 @@
   object Break
   {
     def apply(w: Int): XML.Tree =
-      XML.Elem(
-        Markup(Markup.BREAK, List((Markup.WIDTH, Markup.Int(w)))), List(XML.Text(Symbol.spaces(w))))
+      XML.Elem(Markup(Markup.BREAK, Markup.Width(w)), List(XML.Text(Symbol.spaces(w))))
 
     def unapply(tree: XML.Tree): Option[Int] =
       tree match {
-        case XML.Elem(Markup(Markup.BREAK, List((Markup.WIDTH, Markup.Int(w)))), _) => Some(w)
+        case XML.Elem(Markup(Markup.BREAK, Markup.Width(w)), _) => Some(w)
         case _ => None
       }
   }
--- a/src/Pure/General/scan.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/General/scan.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -105,7 +105,7 @@
 
 fun catch scan xs = scan xs
   handle ABORT msg => raise Fail msg
-    | FAIL msg => raise Fail (the_default "Syntax error." msg);
+    | FAIL msg => raise Fail (the_default "Syntax error" msg);
 
 
 (* scanner combinators *)
--- a/src/Pure/General/secure.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/General/secure.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -13,7 +13,6 @@
   val use_text: use_context -> int * string -> bool -> string -> unit
   val use_file: use_context -> bool -> string -> unit
   val toplevel_pp: string list -> string -> unit
-  val Isar_setup: unit -> unit
   val PG_setup: unit -> unit
   val commit: unit -> unit
   val bash_output: string -> string * int
@@ -56,8 +55,8 @@
 
 fun commit () = use_global "commit();";   (*commit is dynamically bound!*)
 
-fun Isar_setup () = use_global "open Unsynchronized;";
-fun PG_setup () = use_global "structure ThyLoad = ProofGeneral.ThyLoad;";
+fun PG_setup () =
+  use_global "val change = Unsynchronized.change; structure ThyLoad = ProofGeneral.ThyLoad;";
 
 
 (* shell commands *)
--- a/src/Pure/General/symbol.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/General/symbol.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -53,6 +53,9 @@
 
   def is_plain(c: Char): Boolean = !(c == '\r' || c == '\\' || '\ud800' <= c && c <= '\udfff')
 
+  def is_physical_newline(s: CharSequence): Boolean =
+    "\n".contentEquals(s) || "\r".contentEquals(s) || "\r\n".contentEquals(s)
+
   def is_wellformed(s: CharSequence): Boolean =
     s.length == 1 && is_plain(s.charAt(0)) || !bad_symbol.pattern.matcher(s).matches
 
--- a/src/Pure/General/xml.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/General/xml.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -102,17 +102,19 @@
         x
       }
 
+      def trim_bytes(s: String): String = new String(s.toCharArray)
+
       def cache_string(x: String): String =
         lookup(x) match {
           case Some(y) => y
-          case None => store(new String(x.toCharArray))  // trim string value
+          case None => store(trim_bytes(x))
         }
       def cache_props(x: List[(String, String)]): List[(String, String)] =
         if (x.isEmpty) x
         else
           lookup(x) match {
             case Some(y) => y
-            case None => store(x.map(p => (cache_string(p._1), cache_string(p._2))))
+            case None => store(x.map(p => (trim_bytes(p._1).intern, cache_string(p._2))))
           }
       def cache_markup(x: Markup): Markup =
         lookup(x) match {
--- a/src/Pure/IsaMakefile	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/IsaMakefile	Wed Sep 01 17:19:47 2010 +0200
@@ -6,7 +6,7 @@
 
 default: Pure
 images: Pure
-test: RAW Pure-ProofGeneral
+test: RAW
 all: images test
 
 
@@ -256,15 +256,7 @@
 	@./mk
 
 
-## Proof General keywords
-
-Pure-ProofGeneral: Pure $(LOG)/Pure-ProofGeneral.gz
-
-$(LOG)/Pure-ProofGeneral.gz: $(OUT)/Pure ProofGeneral/proof_general_keywords.ML
-	@$(ISABELLE_TOOL) usedir -f proof_general_keywords.ML $(OUT)/Pure ProofGeneral
-
-
 ## clean
 
 clean:
-	@rm -f $(OUT)/Pure $(LOG)/Pure.gz $(OUT)/RAW $(LOG)/RAW.gz $(LOG)/Pure-ProofGeneral.gz
+	@rm -f $(OUT)/Pure $(LOG)/Pure.gz $(OUT)/RAW $(LOG)/RAW.gz
--- a/src/Pure/Isar/class.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/class.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -368,7 +368,7 @@
 fun gen_classrel mk_prop classrel thy =
   let
     fun after_qed results =
-      ProofContext.theory ((fold o fold) AxClass.add_classrel results);
+      ProofContext.background_theory ((fold o fold) AxClass.add_classrel results);
   in
     thy
     |> ProofContext.init_global
@@ -482,7 +482,7 @@
 
 val type_name = sanitize_name o Long_Name.base_name;
 
-fun define_overloaded (c, U) v (b_def, rhs) = Local_Theory.theory_result
+fun define_overloaded (c, U) v (b_def, rhs) = Local_Theory.background_theory_result
     (AxClass.declare_overloaded (c, U) ##>> AxClass.define_overloaded b_def (c, rhs))
   ##> (map_instantiation o apsnd) (filter_out (fn (_, (v', _)) => v' = v))
   ##> Local_Theory.target synchronize_inst_syntax;
@@ -572,7 +572,7 @@
     val (tycos, vs, sort) = (#arities o the_instantiation) lthy;
     val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos;
     fun after_qed' results =
-      Local_Theory.theory (fold (AxClass.add_arity o Thm.varifyT_global) results)
+      Local_Theory.background_theory (fold (AxClass.add_arity o Thm.varifyT_global) results)
       #> after_qed;
   in
     lthy
@@ -608,7 +608,7 @@
     val (tycos, vs, sort) = read_multi_arity thy raw_arities;
     val sorts = map snd vs;
     val arities = maps (fn tyco => Logic.mk_arities (tyco, sorts, sort)) tycos;
-    fun after_qed results = ProofContext.theory
+    fun after_qed results = ProofContext.background_theory
       ((fold o fold) AxClass.add_arity results);
   in
     thy
--- a/src/Pure/Isar/class_declaration.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/class_declaration.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -126,8 +126,8 @@
                 else error ("Type inference imposes additional sort constraint "
                   ^ Syntax.string_of_sort_global thy inferred_sort
                   ^ " of type parameter " ^ Name.aT ^ " of sort "
-                  ^ Syntax.string_of_sort_global thy a_sort ^ ".")
-            | _ => error "Multiple type variables in class specification.";
+                  ^ Syntax.string_of_sort_global thy a_sort)
+            | _ => error "Multiple type variables in class specification";
       in (map o map_atyps) (K (TFree (Name.aT, fixate_sort))) Ts end;
     val after_infer_fixate = (map o map_atyps)
       (fn T as TFree _ => T | T as TVar (vi, sort) =>
@@ -330,7 +330,7 @@
     val some_prop = try the_single props;
     val some_dep_morph = try the_single (map snd deps);
     fun after_qed some_wit =
-      ProofContext.theory (Class.register_subclass (sub, sup)
+      ProofContext.background_theory (Class.register_subclass (sub, sup)
         some_dep_morph some_wit export)
       #> ProofContext.theory_of #> Named_Target.init sub;
   in do_proof after_qed some_prop goal_ctxt end;
--- a/src/Pure/Isar/expression.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/expression.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -824,8 +824,9 @@
     val goal_ctxt = fold Variable.auto_fixes eqns expr_ctxt;
     val export' = Variable.export_morphism goal_ctxt expr_ctxt;
 
-    fun after_qed witss eqns = (ProofContext.theory o Context.theory_map)
-      (note_eqns_register deps witss attrss eqns export export');
+    fun after_qed witss eqns =
+      (ProofContext.background_theory o Context.theory_map)
+        (note_eqns_register deps witss attrss eqns export export');
 
   in Element.witness_proof_eqs after_qed propss eqns goal_ctxt end;
 
@@ -872,7 +873,7 @@
     val target = intern thy raw_target;
     val target_ctxt = Named_Target.init target thy;
     val ((propss, deps, export), goal_ctxt) = prep_expr expression target_ctxt;
-    fun after_qed witss = ProofContext.theory
+    fun after_qed witss = ProofContext.background_theory
       (fold (fn ((dep, morph), wits) => Locale.add_dependency
         target (dep, morph $> Element.satisfy_morphism wits) export) (deps ~~ witss));
   in Element.witness_proof after_qed propss goal_ctxt end;
--- a/src/Pure/Isar/generic_target.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/generic_target.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -41,7 +41,7 @@
 fun check_mixfix ctxt (b, extra_tfrees) mx =
   if null extra_tfrees then mx
   else
-    (warning
+    (Context_Position.if_visible ctxt warning
       ("Additional type variable(s) in specification of " ^ Binding.str_of b ^ ": " ^
         commas (map (Syntax.string_of_typ ctxt o TFree) (sort_wrt #1 extra_tfrees)) ^
         (if mx = NoSyn then ""
@@ -195,16 +195,16 @@
       (Morphism.transform (Local_Theory.global_morphism lthy) decl);
   in
     lthy
-    |> Local_Theory.theory (Context.theory_map global_decl)
+    |> Local_Theory.background_theory (Context.theory_map global_decl)
     |> Local_Theory.target (Context.proof_map global_decl)
   end;
 
 fun theory_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) lthy =
   let
-    val (const, lthy2) = lthy |> Local_Theory.theory_result
+    val (const, lthy2) = lthy |> Local_Theory.background_theory_result
       (Sign.declare_const ((b, U), mx));
     val lhs = list_comb (const, type_params @ term_params);
-    val ((_, def), lthy3) = lthy2 |> Local_Theory.theory_result
+    val ((_, def), lthy3) = lthy2 |> Local_Theory.background_theory_result
       (Thm.add_def false false (b_def, Logic.mk_equals (lhs, rhs)));
   in ((lhs, def), lthy3) end;
 
@@ -214,12 +214,13 @@
     val global_facts' = Attrib.map_facts (Attrib.attribute_i thy) global_facts;
   in
     lthy
-    |> Local_Theory.theory (PureThy.note_thmss kind global_facts' #> snd)
+    |> Local_Theory.background_theory (PureThy.note_thmss kind global_facts' #> snd)
     |> Local_Theory.target (ProofContext.note_thmss kind global_facts' #> snd)
   end;
 
-fun theory_abbrev prmode ((b, mx), t) = Local_Theory.theory
-  (Sign.add_abbrev (#1 prmode) (b, t) #->
-    (fn (lhs, _) => Sign.notation true prmode [(lhs, mx)]));
+fun theory_abbrev prmode ((b, mx), t) =
+  Local_Theory.background_theory
+    (Sign.add_abbrev (#1 prmode) (b, t) #->
+      (fn (lhs, _) => Sign.notation true prmode [(lhs, mx)]));
 
 end;
--- a/src/Pure/Isar/isar_cmd.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/isar_cmd.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -44,7 +44,6 @@
   val pwd: Toplevel.transition -> Toplevel.transition
   val display_drafts: Path.T list -> Toplevel.transition -> Toplevel.transition
   val print_drafts: Path.T list -> Toplevel.transition -> Toplevel.transition
-  val pretty_setmargin: int -> Toplevel.transition -> Toplevel.transition
   val print_context: Toplevel.transition -> Toplevel.transition
   val print_theory: bool -> Toplevel.transition -> Toplevel.transition
   val print_syntax: Toplevel.transition -> Toplevel.transition
@@ -321,11 +320,6 @@
   in File.isabelle_tool "print" ("-c " ^ outfile); () end);
 
 
-(* pretty_setmargin *)
-
-fun pretty_setmargin n = Toplevel.imperative (fn () => Pretty.margin_default := n);
-
-
 (* print parts of theory and proof context *)
 
 val print_context = Toplevel.keep Toplevel.print_state_context;
--- a/src/Pure/Isar/isar_syn.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/isar_syn.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -786,7 +786,8 @@
 
 val _ =
   Outer_Syntax.improper_command "pretty_setmargin" "change default margin for pretty printing"
-    Keyword.diag (Parse.nat >> (Toplevel.no_timing oo Isar_Cmd.pretty_setmargin));
+    Keyword.diag (Parse.nat >>
+      (fn n => Toplevel.no_timing o Toplevel.imperative (fn () => Pretty.margin_default := n)));
 
 val _ =
   Outer_Syntax.improper_command "help" "print outer syntax commands" Keyword.diag
@@ -991,7 +992,7 @@
     (Scan.succeed
       (Toplevel.no_timing o Toplevel.keep (fn state =>
         (Context.set_thread_data (try Toplevel.generic_theory_of state);
-          raise Toplevel.TERMINATE))));
+          raise Runtime.TERMINATE))));
 
 end;
 
--- a/src/Pure/Isar/local_theory.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/local_theory.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -21,8 +21,8 @@
   val target_of: local_theory -> Proof.context
   val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
   val raw_theory: (theory -> theory) -> local_theory -> local_theory
-  val theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
-  val theory: (theory -> theory) -> local_theory -> local_theory
+  val background_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
+  val background_theory: (theory -> theory) -> local_theory -> local_theory
   val target_result: (Proof.context -> 'a * Proof.context) -> local_theory -> 'a * local_theory
   val target: (Proof.context -> Proof.context) -> local_theory -> local_theory
   val map_contexts: (Context.generic -> Context.generic) -> local_theory -> local_theory
@@ -144,7 +144,7 @@
 
 val checkpoint = raw_theory Theory.checkpoint;
 
-fun theory_result f lthy =
+fun background_theory_result f lthy =
   lthy |> raw_theory_result (fn thy =>
     thy
     |> Sign.map_naming (K (naming_of lthy))
@@ -152,7 +152,7 @@
     ||> Sign.restore_naming thy
     ||> Theory.checkpoint);
 
-fun theory f = #2 o theory_result (f #> pair ());
+fun background_theory f = #2 o background_theory_result (f #> pair ());
 
 fun target_result f lthy =
   let
@@ -169,7 +169,7 @@
 fun target f = #2 o target_result (f #> pair ());
 
 fun map_contexts f =
-  theory (Context.theory_map f) #>
+  background_theory (Context.theory_map f) #>
   target (Context.proof_map f) #>
   Context.proof_map f;
 
@@ -261,7 +261,7 @@
 
 (* exit *)
 
-val exit = ProofContext.theory Theory.checkpoint o operation #exit;
+val exit = ProofContext.background_theory Theory.checkpoint o operation #exit;
 val exit_global = ProofContext.theory_of o exit;
 
 fun exit_result f (x, lthy) =
--- a/src/Pure/Isar/locale.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/locale.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -483,7 +483,7 @@
     val thy' = (change_locale loc o apsnd) (cons ((name, (morph, export)), serial ())) thy;
     val context' = Context.Theory thy';
     val (_, regs) = fold_rev (roundup thy' cons export)
-      (all_registrations context') (get_idents (context'), []);
+      (registrations_of context' loc) (get_idents (context'), []);
   in
     thy'
     |> fold_rev (fn dep => Context.theory_map (add_registration dep NONE export)) regs
@@ -497,8 +497,8 @@
 fun add_thmss loc kind args ctxt =
   let
     val (Notes args', ctxt') = Element.activate_i (Notes (kind, args)) ctxt;
-    val ctxt'' = ctxt' |> ProofContext.theory (
-      (change_locale loc o apfst o apsnd) (cons (args', serial ()))
+    val ctxt'' = ctxt' |> ProofContext.background_theory
+     ((change_locale loc o apfst o apsnd) (cons (args', serial ()))
         #>
       (* Registrations *)
       (fn thy => fold_rev (fn (_, morph) =>
@@ -519,7 +519,7 @@
       [([Drule.dummy_thm], [])])];
 
 fun add_syntax_declaration loc decl =
-  ProofContext.theory ((change_locale loc o apfst o apfst) (cons (decl, serial ())))
+  ProofContext.background_theory ((change_locale loc o apfst o apfst) (cons (decl, serial ())))
   #> add_declaration loc decl;
 
 
--- a/src/Pure/Isar/named_target.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/named_target.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -118,7 +118,7 @@
       (Element.morph_ctxt (Local_Theory.target_morphism lthy)) local_facts;
   in
     lthy
-    |> Local_Theory.theory (PureThy.note_thmss kind global_facts' #> snd)
+    |> Local_Theory.background_theory (PureThy.note_thmss kind global_facts' #> snd)
     |> Local_Theory.target (Locale.add_thmss locale kind local_facts')
   end
 
@@ -129,19 +129,21 @@
 
 (* abbrev *)
 
-fun locale_abbrev ta prmode ((b, mx), t) xs = Local_Theory.theory_result
-  (Sign.add_abbrev Print_Mode.internal (b, t)) #->
-    (fn (lhs, _) => locale_const_declaration ta prmode
-      ((b, mx), Term.list_comb (Logic.unvarify_global lhs, xs)));
+fun locale_abbrev ta prmode ((b, mx), t) xs =
+  Local_Theory.background_theory_result
+    (Sign.add_abbrev Print_Mode.internal (b, t)) #->
+      (fn (lhs, _) => locale_const_declaration ta prmode
+        ((b, mx), Term.list_comb (Logic.unvarify_global lhs, xs)));
 
 fun target_abbrev (ta as Target {target, is_locale, is_class, ...})
     prmode (b, mx) (global_rhs, t') xs lthy =
-  if is_locale
-    then lthy
-      |> locale_abbrev ta prmode ((b, if is_class then NoSyn else mx), global_rhs) xs
-      |> is_class ? Class.abbrev target prmode ((b, mx), t')
-    else lthy
-      |> Generic_Target.theory_abbrev prmode ((b, mx), global_rhs);
+  if is_locale then
+    lthy
+    |> locale_abbrev ta prmode ((b, if is_class then NoSyn else mx), global_rhs) xs
+    |> is_class ? Class.abbrev target prmode ((b, mx), t')
+  else
+    lthy
+    |> Generic_Target.theory_abbrev prmode ((b, mx), global_rhs);
 
 
 (* pretty *)
@@ -200,9 +202,10 @@
 
 fun init target thy = init_target (named_target thy target) thy;
 
-fun reinit lthy = case peek lthy
- of SOME {target, ...} => init target o Local_Theory.exit_global
-  | NONE => error "Not in a named target";
+fun reinit lthy =
+  (case peek lthy of
+    SOME {target, ...} => init target o Local_Theory.exit_global
+  | NONE => error "Not in a named target");
 
 val theory_init = init_target global_target;
 
--- a/src/Pure/Isar/obtain.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/obtain.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -181,7 +181,7 @@
       if Thm.concl_of th aconv thesis andalso
         Logic.strip_assums_concl prem aconv thesis then th
       else error ("Guessed a different clause:\n" ^ Display.string_of_thm ctxt th)
-  | [] => error "Goal solved -- nothing guessed."
+  | [] => error "Goal solved -- nothing guessed"
   | _ => error ("Guess split into several cases:\n" ^ Display.string_of_thm ctxt th));
 
 fun result tac facts ctxt =
--- a/src/Pure/Isar/overloading.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/overloading.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -140,7 +140,7 @@
   end
 
 fun define_overloaded (c, U) (v, checked) (b_def, rhs) =
-  Local_Theory.theory_result
+  Local_Theory.background_theory_result
     (Thm.add_def (not checked) true (b_def, Logic.mk_equals (Const (c, Term.fastype_of rhs), rhs)))
   ##> map_overloading (filter_out (fn (_, (v', _)) => v' = v))
   ##> Local_Theory.target synchronize_syntax
--- a/src/Pure/Isar/proof.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/proof.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -42,6 +42,7 @@
   val raw_goal: state -> {context: context, facts: thm list, goal: thm}
   val goal: state -> {context: context, facts: thm list, goal: thm}
   val simple_goal: state -> {context: context, goal: thm}
+  val status_markup: state -> Markup.T
   val let_bind: (term list * term) list -> state -> state
   val let_bind_cmd: (string list * string) list -> state -> state
   val write: Syntax.mode -> (term * mixfix) list -> state -> state
@@ -520,6 +521,11 @@
     val (ctxt, (_, goal)) = get_goal (refine_insert facts state);
   in {context = ctxt, goal = goal} end;
 
+fun status_markup state =
+  (case try goal state of
+    SOME {goal, ...} => Markup.proof_state (Thm.nprems_of goal)
+  | NONE => Markup.empty);
+
 
 
 (*** structured proof commands ***)
--- a/src/Pure/Isar/proof_context.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/proof_context.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -38,8 +38,8 @@
   val cases_of: Proof.context -> (string * (Rule_Cases.T * bool)) list
   val transfer_syntax: theory -> Proof.context -> Proof.context
   val transfer: theory -> Proof.context -> Proof.context
-  val theory: (theory -> theory) -> Proof.context -> Proof.context
-  val theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context
+  val background_theory: (theory -> theory) -> Proof.context -> Proof.context
+  val background_theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context
   val extern_fact: Proof.context -> string -> xstring
   val pretty_term_abbrev: Proof.context -> term -> Pretty.T
   val pretty_fact_aux: Proof.context -> bool -> string * thm list -> Pretty.T
@@ -283,9 +283,9 @@
 
 fun transfer thy = Context.raw_transfer thy #> transfer_syntax thy;
 
-fun theory f ctxt = transfer (f (theory_of ctxt)) ctxt;
+fun background_theory f ctxt = transfer (f (theory_of ctxt)) ctxt;
 
-fun theory_result f ctxt =
+fun background_theory_result f ctxt =
   let val (res, thy') = f (theory_of ctxt)
   in (res, ctxt |> transfer thy') end;
 
--- a/src/Pure/Isar/runtime.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/runtime.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -11,6 +11,7 @@
   exception EXCURSION_FAIL of exn * string
   exception TOPLEVEL_ERROR
   val exn_context: Proof.context option -> exn -> exn
+  val exn_messages: (exn -> Position.T) -> exn -> string list
   val exn_message: (exn -> Position.T) -> exn -> string
   val debugging: ('a -> 'b) -> 'a -> 'b
   val controlled_execution: ('a -> 'b) -> 'a -> 'b
@@ -41,7 +42,7 @@
 fun if_context NONE _ _ = []
   | if_context (SOME ctxt) f xs = map (f ctxt) xs;
 
-fun exn_message exn_position e =
+fun exn_messages exn_position e =
   let
     fun raised exn name msgs =
       let val pos = Position.str_of (exn_position exn) in
@@ -53,32 +54,36 @@
 
     val detailed = ! Output.debugging;
 
-    fun exn_msg _ (CONTEXT (ctxt, exn)) = exn_msg (SOME ctxt) exn
-      | exn_msg _ (Exn.EXCEPTIONS []) = "Exception."
-      | exn_msg ctxt (Exn.EXCEPTIONS exns) = cat_lines (map (exn_msg ctxt) exns)
-      | exn_msg ctxt (EXCURSION_FAIL (exn, loc)) =
-          exn_msg ctxt exn ^ Markup.markup Markup.location ("\n" ^ loc)
-      | exn_msg _ TERMINATE = "Exit."
-      | exn_msg _ Exn.Interrupt = "Interrupt."
-      | exn_msg _ TimeLimit.TimeOut = "Timeout."
-      | exn_msg _ TOPLEVEL_ERROR = "Error."
-      | exn_msg _ (SYS_ERROR msg) = "## SYSTEM ERROR ##\n" ^ msg
-      | exn_msg _ (ERROR msg) = msg
-      | exn_msg _ (exn as Fail msg) = raised exn "Fail" [msg]
-      | exn_msg _ (exn as THEORY (msg, thys)) =
-          raised exn "THEORY" (msg :: (if detailed then map Context.str_of_thy thys else []))
-      | exn_msg _ (exn as Syntax.AST (msg, asts)) = raised exn "AST" (msg ::
-            (if detailed then map (Pretty.string_of o Syntax.pretty_ast) asts else []))
-      | exn_msg ctxt (exn as TYPE (msg, Ts, ts)) = raised exn "TYPE" (msg ::
+    fun exn_msgs _ (CONTEXT (ctxt, exn)) = exn_msgs (SOME ctxt) exn
+      | exn_msgs ctxt (Exn.EXCEPTIONS exns) = maps (exn_msgs ctxt) exns
+      | exn_msgs ctxt (EXCURSION_FAIL (exn, loc)) =
+          map (fn msg => msg ^ Markup.markup Markup.location ("\n" ^ loc)) (exn_msgs ctxt exn)
+      | exn_msgs _ TERMINATE = ["Exit"]
+      | exn_msgs _ Exn.Interrupt = []
+      | exn_msgs _ TimeLimit.TimeOut = ["Timeout"]
+      | exn_msgs _ TOPLEVEL_ERROR = ["Error"]
+      | exn_msgs _ (SYS_ERROR msg) = ["## SYSTEM ERROR ##\n" ^ msg]
+      | exn_msgs _ (ERROR msg) = [msg]
+      | exn_msgs _ (exn as Fail msg) = [raised exn "Fail" [msg]]
+      | exn_msgs _ (exn as THEORY (msg, thys)) =
+          [raised exn "THEORY" (msg :: (if detailed then map Context.str_of_thy thys else []))]
+      | exn_msgs _ (exn as Syntax.AST (msg, asts)) = [raised exn "AST" (msg ::
+            (if detailed then map (Pretty.string_of o Syntax.pretty_ast) asts else []))]
+      | exn_msgs ctxt (exn as TYPE (msg, Ts, ts)) = [raised exn "TYPE" (msg ::
             (if detailed then
               if_context ctxt Syntax.string_of_typ Ts @ if_context ctxt Syntax.string_of_term ts
-             else []))
-      | exn_msg ctxt (exn as TERM (msg, ts)) = raised exn "TERM" (msg ::
-            (if detailed then if_context ctxt Syntax.string_of_term ts else []))
-      | exn_msg ctxt (exn as THM (msg, i, thms)) = raised exn ("THM " ^ string_of_int i) (msg ::
-            (if detailed then if_context ctxt Display.string_of_thm thms else []))
-      | exn_msg _ exn = raised exn (General.exnMessage exn) [];
-  in exn_msg NONE e end;
+             else []))]
+      | exn_msgs ctxt (exn as TERM (msg, ts)) = [raised exn "TERM" (msg ::
+            (if detailed then if_context ctxt Syntax.string_of_term ts else []))]
+      | exn_msgs ctxt (exn as THM (msg, i, thms)) = [raised exn ("THM " ^ string_of_int i) (msg ::
+            (if detailed then if_context ctxt Display.string_of_thm thms else []))]
+      | exn_msgs _ exn = [raised exn (General.exnMessage exn) []];
+  in exn_msgs NONE e end;
+
+fun exn_message exn_position exn =
+  (case exn_messages exn_position exn of
+    [] => "Interrupt"
+  | msgs => cat_lines msgs);
 
 
 
--- a/src/Pure/Isar/toplevel.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/toplevel.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -30,19 +30,20 @@
   val timing: bool Unsynchronized.ref
   val profiling: int Unsynchronized.ref
   val skip_proofs: bool Unsynchronized.ref
-  exception TERMINATE
   exception TOPLEVEL_ERROR
   val program: (unit -> 'a) -> 'a
   val thread: bool -> (unit -> unit) -> Thread.thread
   type transition
   val empty: transition
   val init_of: transition -> string option
+  val print_of: transition -> bool
   val name_of: transition -> string
   val pos_of: transition -> Position.T
   val str_of: transition -> string
   val name: string -> transition -> transition
   val position: Position.T -> transition -> transition
   val interactive: bool -> transition -> transition
+  val set_print: bool -> transition -> transition
   val print: transition -> transition
   val no_timing: transition -> transition
   val init_theory: string -> (unit -> theory) -> transition -> transition
@@ -84,10 +85,9 @@
   val unknown_context: transition -> transition
   val setmp_thread_position: transition -> ('a -> 'b) -> 'a -> 'b
   val status: transition -> Markup.T -> unit
-  val error_msg: transition -> exn * string -> unit
+  val error_msg: transition -> string -> unit
   val add_hook: (transition -> state -> state -> unit) -> unit
   val transition: bool -> transition -> state -> (state * (exn * string) option) option
-  val run_command: string -> transition -> state -> state option
   val command: transition -> state -> state
   val excursion: (transition * transition list) list -> (transition * state) list lazy * theory
 end;
@@ -222,8 +222,6 @@
 val profiling = Unsynchronized.ref 0;
 val skip_proofs = Unsynchronized.ref false;
 
-exception TERMINATE = Runtime.TERMINATE;
-exception EXCURSION_FAIL = Runtime.EXCURSION_FAIL;
 exception TOPLEVEL_ERROR = Runtime.TOPLEVEL_ERROR;
 
 fun program body =
@@ -351,12 +349,13 @@
 fun init_of (Transition {trans = [Init (name, _)], ...}) = SOME name
   | init_of _ = NONE;
 
+fun print_of (Transition {print, ...}) = print;
 fun name_of (Transition {name, ...}) = name;
 fun pos_of (Transition {pos, ...}) = pos;
 fun str_of tr = quote (name_of tr) ^ Position.str_of (pos_of tr);
 
 fun command_msg msg tr = msg ^ "command " ^ str_of tr;
-fun at_command tr = command_msg "At " tr ^ ".";
+fun at_command tr = command_msg "At " tr;
 
 fun type_error tr state =
   ERROR (command_msg "Illegal application of " tr ^ " " ^ str_of_state state);
@@ -563,16 +562,8 @@
 fun status tr m =
   setmp_thread_position tr (fn () => Output.status (Markup.markup m "")) ();
 
-fun async_state (tr as Transition {print, ...}) st =
-  if print then
-    ignore
-      (Future.fork (fn () =>
-          setmp_thread_position tr (fn () => Future.status (fn () => print_state false st)) ()))
-  else ();
-
-fun error_msg tr exn_info =
-  setmp_thread_position tr (fn () =>
-    Output.error_msg (ML_Compiler.exn_message (EXCURSION_FAIL exn_info))) ();
+fun error_msg tr msg =
+  setmp_thread_position tr (fn () => Output.error_msg msg) ();
 
 
 (* post-transition hooks *)
@@ -616,8 +607,8 @@
     val ctxt = try context_of st;
     val res =
       (case app int tr st of
-        (_, SOME TERMINATE) => NONE
-      | (st', SOME (EXCURSION_FAIL exn_info)) => SOME (st', SOME exn_info)
+        (_, SOME Runtime.TERMINATE) => NONE
+      | (st', SOME (Runtime.EXCURSION_FAIL exn_info)) => SOME (st', SOME exn_info)
       | (st', SOME exn) => SOME (st', SOME (Runtime.exn_context ctxt exn, at_command tr))
       | (st', NONE) => SOME (st', NONE));
     val _ = (case res of SOME (st', NONE) => apply_hooks st' | _ => ());
@@ -626,32 +617,13 @@
 end;
 
 
-(* managed execution *)
-
-fun run_command thy_name tr st =
-  (case
-      (case init_of tr of
-        SOME name => Exn.capture (fn () => Thy_Header.consistent_name thy_name name) ()
-      | NONE => Exn.Result ()) of
-    Exn.Result () =>
-      let val int = is_some (init_of tr) in
-        (case transition int tr st of
-          SOME (st', NONE) =>
-            (status tr Markup.finished; if int then () else async_state tr st'; SOME st')
-        | SOME (_, SOME (exn as Exn.Interrupt, _)) => reraise exn
-        | SOME (_, SOME exn_info) => (error_msg tr exn_info; status tr Markup.failed; NONE)
-        | NONE => (error_msg tr (TERMINATE, at_command tr); status tr Markup.failed; NONE))
-      end
-  | Exn.Exn exn => (error_msg tr (exn, at_command tr); status tr Markup.failed; NONE));
-
-
 (* nested commands *)
 
 fun command tr st =
   (case transition (! interact) tr st of
     SOME (st', NONE) => st'
-  | SOME (_, SOME exn_info) => raise EXCURSION_FAIL exn_info
-  | NONE => raise EXCURSION_FAIL (TERMINATE, at_command tr));
+  | SOME (_, SOME exn_info) => raise Runtime.EXCURSION_FAIL exn_info
+  | NONE => raise Runtime.EXCURSION_FAIL (Runtime.TERMINATE, at_command tr));
 
 fun command_result tr st =
   let val st' = command tr st
--- a/src/Pure/Isar/typedecl.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Isar/typedecl.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -34,7 +34,7 @@
 fun basic_decl decl (b, n, mx) lthy =
   let val name = Local_Theory.full_name lthy b in
     lthy
-    |> Local_Theory.theory (decl name)
+    |> Local_Theory.background_theory (decl name)
     |> Local_Theory.type_notation true Syntax.mode_default [(Type (name, replicate n dummyT), mx)]
     |> Local_Theory.type_alias b name
     |> pair name
@@ -102,9 +102,10 @@
     val ignored = Term.fold_atyps_sorts (fn (_, []) => I | (T, _) => insert (op =) T) rhs [];
     val _ =
       if null ignored then ()
-      else warning ("Ignoring sort constraints in type variables(s): " ^
-        commas_quote (map (Syntax.string_of_typ ctxt) (rev ignored)) ^
-        "\nin type abbreviation " ^ quote (Binding.str_of b));
+      else Context_Position.if_visible ctxt warning
+        ("Ignoring sort constraints in type variables(s): " ^
+          commas_quote (map (Syntax.string_of_typ ctxt) (rev ignored)) ^
+          "\nin type abbreviation " ^ quote (Binding.str_of b));
   in rhs end;
 
 in
--- a/src/Pure/ML-Systems/unsynchronized.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/ML-Systems/unsynchronized.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -12,10 +12,6 @@
 val op := = op :=;
 val ! = !;
 
-fun set flag = (flag := true; true);
-fun reset flag = (flag := false; false);
-fun toggle flag = (flag := not (! flag); ! flag);
-
 fun change r f = r := f (! r);
 fun change_result r f = let val (x, y) = f (! r) in r := y; x end;
 
--- a/src/Pure/ML/ml_compiler.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/ML/ml_compiler.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -7,6 +7,7 @@
 signature ML_COMPILER =
 sig
   val exn_position: exn -> Position.T
+  val exn_messages: exn -> string list
   val exn_message: exn -> string
   val eval: bool -> Position.T -> ML_Lex.token list -> unit
 end
@@ -15,6 +16,7 @@
 struct
 
 fun exn_position _ = Position.none;
+val exn_messages = Runtime.exn_messages exn_position;
 val exn_message = Runtime.exn_message exn_position;
 
 fun eval verbose pos toks =
--- a/src/Pure/ML/ml_compiler_polyml-5.3.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/ML/ml_compiler_polyml-5.3.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -4,13 +4,6 @@
 Advanced runtime compilation for Poly/ML 5.3.0.
 *)
 
-signature ML_COMPILER =
-sig
-  val exn_position: exn -> Position.T
-  val exn_message: exn -> string
-  val eval: bool -> Position.T -> ML_Lex.token list -> unit
-end
-
 structure ML_Compiler: ML_COMPILER =
 struct
 
@@ -37,6 +30,7 @@
     NONE => Position.none
   | SOME loc => position_of loc);
 
+val exn_messages = Runtime.exn_messages exn_position;
 val exn_message = Runtime.exn_message exn_position;
 
 
@@ -44,13 +38,18 @@
 
 fun report_parse_tree depth space =
   let
+    val report_text =
+      (case Context.thread_data () of
+        SOME (Context.Proof ctxt) => Context_Position.report_text ctxt
+      | _ => Position.report_text);
+
     fun report_decl markup loc decl =
-      Position.report_text Markup.ML_ref (position_of loc)
+      report_text Markup.ML_ref (position_of loc)
         (Markup.markup (Markup.properties (Position.properties_of (position_of decl)) markup) "");
     fun report loc (PolyML.PTtype types) =
           PolyML.NameSpace.displayTypeExpression (types, depth, space)
           |> pretty_ml |> Pretty.from_ML |> Pretty.string_of
-          |> Position.report_text Markup.ML_typing (position_of loc)
+          |> report_text Markup.ML_typing (position_of loc)
       | report loc (PolyML.PTdeclaredAt decl) = report_decl Markup.ML_def loc decl
       | report loc (PolyML.PTopenedAt decl) = report_decl Markup.ML_open loc decl
       | report loc (PolyML.PTstructureAt decl) = report_decl Markup.ML_struct loc decl
--- a/src/Pure/ML/ml_context.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/ML/ml_context.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -35,8 +35,8 @@
   val eval_text_in: Proof.context option -> bool -> Position.T -> Symbol_Pos.text -> unit
   val expression: Position.T -> string -> string -> ML_Lex.token Antiquote.antiquote list ->
     Context.generic -> Context.generic
-  val evaluate:
-    Proof.context -> bool -> string * (unit -> 'a) option Unsynchronized.ref -> string -> 'a
+  val evaluate: Proof.context -> bool ->
+    string * (unit -> 'a) option Unsynchronized.ref -> string * string -> 'a
 end
 
 structure ML_Context: ML_CONTEXT =
@@ -166,7 +166,9 @@
 fun eval verbose pos ants =
   let
     (*prepare source text*)
-    val ((env, body), env_ctxt) = eval_antiquotes (ants, pos) (Context.thread_data ());
+    val ((env, body), env_ctxt) =
+      eval_antiquotes (ants, pos) (Context.thread_data ())
+      ||> Option.map (Context.mapping I (Context_Position.set_visible false));
     val _ =
       if ! trace then tracing (cat_lines [ML_Lex.flatten env, ML_Lex.flatten body])
       else ();
@@ -201,10 +203,10 @@
 
 
 (* FIXME not thread-safe -- reference can be accessed directly *)
-fun evaluate ctxt verbose (ref_name, r) txt = NAMED_CRITICAL "ML" (fn () =>
+fun evaluate ctxt verbose (ref_name, r) (prelude, txt) = NAMED_CRITICAL "ML" (fn () =>
   let
-    val ants =
-      ML_Lex.read Position.none ("val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))");
+    val s = prelude ^ "val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))"
+    val ants = ML_Lex.read Position.none s;
     val _ = r := NONE;
     val _ =
       Context.setmp_thread_data (SOME (Context.Proof ctxt))
--- a/src/Pure/PIDE/command.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/PIDE/command.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -8,23 +8,25 @@
 package isabelle
 
 
+import scala.collection.immutable.SortedMap
+
+
 object Command
 {
   /** accumulated results from prover **/
 
   case class State(
     val command: Command,
-    val status: List[Markup],
-    val reverse_results: List[XML.Tree],
-    val markup: Markup_Tree)
+    val status: List[Markup] = Nil,
+    val results: SortedMap[Long, XML.Tree] = SortedMap.empty,
+    val markup: Markup_Tree = Markup_Tree.empty)
   {
     /* content */
 
-    lazy val results = reverse_results.reverse
-
     def add_status(st: Markup): State = copy(status = st :: status)
     def add_markup(info: Text.Info[Any]): State = copy(markup = markup + info)
-    def add_result(result: XML.Tree): State = copy(reverse_results = result :: reverse_results)
+    def add_result(serial: Long, result: XML.Tree): State =
+      copy(results = results + (serial -> result))
 
     def root_info: Text.Info[Any] =
       new Text.Info(command.range,
@@ -34,7 +36,7 @@
 
     /* message dispatch */
 
-    def accumulate(message: XML.Tree): Command.State =
+    def accumulate(message: XML.Elem): Command.State =
       message match {
         case XML.Elem(Markup(Markup.STATUS, _), msgs) =>
           (this /: msgs)((state, msg) =>
@@ -46,15 +48,22 @@
         case XML.Elem(Markup(Markup.REPORT, _), msgs) =>
           (this /: msgs)((state, msg) =>
             msg match {
-              case XML.Elem(Markup(name, atts), args)
-              if Position.get_id(atts) == Some(command.id) && Position.get_range(atts).isDefined =>
-                val range = command.decode(Position.get_range(atts).get)
-                val props = atts.filterNot(p => Markup.POSITION_PROPERTIES(p._1))
-                val info = Text.Info[Any](range, XML.Elem(Markup(name, props), args))
+              case XML.Elem(Markup(name, atts @ Position.Id_Range(id, range)), args)
+              if id == command.id =>
+                val props = Position.purge(atts)
+                val info =
+                  Text.Info[Any](command.decode(range), XML.Elem(Markup(name, props), args))
                 state.add_markup(info)
               case _ => System.err.println("Ignored report message: " + msg); state
             })
-        case _ => add_result(message)
+        case XML.Elem(Markup(name, atts), body) =>
+          atts match {
+            case Markup.Serial(i) =>
+              val result = XML.Elem(Markup(name, Position.purge(atts)), body)
+              (add_result(i, result) /: Isar_Document.reported_positions(command.id, message))(
+                (st, range) => st.add_markup(Text.Info(command.decode(range), result)))
+            case _ => System.err.println("Ignored message without serial number: " + message); this
+          }
       }
   }
 
@@ -89,6 +98,10 @@
   def source(range: Text.Range): String = source.substring(range.start, range.stop)
   def length: Int = source.length
 
+  val newlines =
+    (0 /: Symbol.iterator(source)) {
+      case (n, s) => if (Symbol.is_physical_newline(s)) n + 1 else n }
+
   val range: Text.Range = Text.Range(0, length)
 
   lazy val symbol_index = new Symbol.Index(source)
@@ -98,5 +111,5 @@
 
   /* accumulated results */
 
-  val empty_state: Command.State = Command.State(this, Nil, Nil, Markup_Tree.empty)
+  val empty_state: Command.State = Command.State(this)
 }
--- a/src/Pure/PIDE/document.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/PIDE/document.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -192,6 +192,59 @@
 
 
 
+(* toplevel transactions *)
+
+local
+
+fun proof_status tr st =
+  (case try Toplevel.proof_of st of
+    SOME prf => Toplevel.status tr (Proof.status_markup prf)
+  | NONE => ());
+
+fun async_state tr st =
+  if Toplevel.print_of tr then
+    ignore
+      (Future.fork
+        (fn () =>
+          Toplevel.setmp_thread_position tr
+            (fn () => Future.status (fn () => Toplevel.print_state false st)) ()))
+  else ();
+
+in
+
+fun run_command thy_name tr st =
+  (case
+      (case Toplevel.init_of tr of
+        SOME name => Exn.capture (fn () => Thy_Header.consistent_name thy_name name) ()
+      | NONE => Exn.Result ()) of
+    Exn.Result () =>
+      let
+        val int = is_some (Toplevel.init_of tr);
+        val (errs, result) =
+          (case Toplevel.transition int tr st of
+            SOME (st', NONE) => ([], SOME st')
+          | SOME (_, SOME exn_info) =>
+              (case ML_Compiler.exn_messages (Runtime.EXCURSION_FAIL exn_info) of
+                [] => raise Exn.Interrupt
+              | errs => (errs, NONE))
+          | NONE => ([ML_Compiler.exn_message Runtime.TERMINATE], NONE));
+        val _ = List.app (Toplevel.error_msg tr) errs;
+        val _ =
+          (case result of
+            NONE => Toplevel.status tr Markup.failed
+          | SOME st' =>
+             (Toplevel.status tr Markup.finished;
+              proof_status tr st';
+              if int then () else async_state tr st'));
+      in result end
+  | Exn.Exn exn =>
+     (Toplevel.error_msg tr (ML_Compiler.exn_message exn); Toplevel.status tr Markup.failed; NONE))
+
+end;
+
+
+
+
 (** editing **)
 
 (* edit *)
@@ -214,7 +267,7 @@
           NONE => NONE
         | SOME st =>
             let val exec_tr = Toplevel.put_id (print_id exec_id') (Future.join future_tr)
-            in Toplevel.run_command name exec_tr st end));
+            in run_command name exec_tr st end));
     val state' = define_exec exec_id' exec' state;
   in (exec_id', (id, exec_id') :: updates, state') end;
 
@@ -263,7 +316,8 @@
         | force_exec (SOME exec_id) = ignore (Lazy.force (the_exec state exec_id));
 
       val _ = List.app Future.cancel execution;
-      fun await_cancellation () = uninterruptible (fn _ => Future.join_results) execution;
+      fun await_cancellation () =
+        uninterruptible (fn _ => fn () => Future.join_results execution) ();
 
       val execution' = (* FIXME proper node deps *)
         node_names_of version |> map (fn name =>
--- a/src/Pure/PIDE/document.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/PIDE/document.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -44,7 +44,6 @@
   {
     val empty: Node = new Node(Linear_Set())
 
-    // FIXME not scalable
     def command_starts(commands: Iterator[Command], offset: Text.Offset = 0)
       : Iterator[(Command, Text.Offset)] =
     {
@@ -57,16 +56,36 @@
     }
   }
 
+  private val block_size = 1024
+
   class Node(val commands: Linear_Set[Command])
   {
-    def command_starts: Iterator[(Command, Text.Offset)] =
-      Node.command_starts(commands.iterator)
+    private lazy val full_index: (Array[(Command, Text.Offset)], Text.Range) =
+    {
+      val blocks = new mutable.ListBuffer[(Command, Text.Offset)]
+      var next_block = 0
+      var last_stop = 0
+      for ((command, start) <- Node.command_starts(commands.iterator)) {
+        last_stop = start + command.length
+        while (last_stop + 1 > next_block) {
+          blocks += (command -> start)
+          next_block += block_size
+        }
+      }
+      (blocks.toArray, Text.Range(0, last_stop))
+    }
 
-    def command_start(cmd: Command): Option[Text.Offset] =
-      command_starts.find(_._1 == cmd).map(_._2)
+    def full_range: Text.Range = full_index._2
 
     def command_range(i: Text.Offset = 0): Iterator[(Command, Text.Offset)] =
-      command_starts dropWhile { case (cmd, start) => start + cmd.length <= i }
+    {
+      if (!commands.isEmpty && full_range.contains(i)) {
+        val (cmd0, start0) = full_index._1(i / block_size)
+        Node.command_starts(commands.iterator(cmd0), start0) dropWhile {
+          case (cmd, start) => start + cmd.length <= i }
+      }
+      else Iterator.empty
+    }
 
     def command_range(range: Text.Range): Iterator[(Command, Text.Offset)] =
       command_range(range.start) takeWhile { case (_, start) => start < range.stop }
@@ -83,6 +102,12 @@
           commands.reverse_iterator(command).find(cmd => !cmd.is_ignored)
         case None => None
       }
+
+    def command_start(cmd: Command): Option[Text.Offset] =
+      command_starts.find(_._1 == cmd).map(_._2)
+
+    def command_starts: Iterator[(Command, Text.Offset)] =
+      Node.command_starts(commands.iterator)
   }
 
 
@@ -116,7 +141,25 @@
   }
 
 
-  /* history navigation and persistent snapshots */
+  /* history navigation */
+
+  object History
+  {
+    val init = new History(List(Change.init))
+  }
+
+  // FIXME pruning, purging of state
+  class History(val undo_list: List[Change])
+  {
+    require(!undo_list.isEmpty)
+
+    def tip: Change = undo_list.head
+    def +(change: Change): History = new History(change :: undo_list)
+  }
+
+
+
+  /** global state -- document structure, execution process, editing history **/
 
   abstract class Snapshot
   {
@@ -129,59 +172,15 @@
     def convert(range: Text.Range): Text.Range = range.map(convert(_))
     def revert(i: Text.Offset): Text.Offset
     def revert(range: Text.Range): Text.Range = range.map(revert(_))
-  }
-
-  object History
-  {
-    val init = new History(List(Change.init))
+    def select_markup[A](range: Text.Range)
+      (result: PartialFunction[Text.Info[Any], A])(default: A): Iterator[Text.Info[A]]
   }
 
-  // FIXME pruning, purging of state
-  class History(undo_list: List[Change])
-  {
-    require(!undo_list.isEmpty)
-
-    def tip: Change = undo_list.head
-    def +(ch: Change): History = new History(ch :: undo_list)
-
-    def snapshot(name: String, pending_edits: List[Text.Edit], state_snapshot: State): Snapshot =
-    {
-      val found_stable = undo_list.find(change =>
-        change.is_finished && state_snapshot.is_assigned(change.current.join))
-      require(found_stable.isDefined)
-      val stable = found_stable.get
-      val latest = undo_list.head
-
-      val edits =
-        (pending_edits /: undo_list.takeWhile(_ != stable))((edits, change) =>
-            (for ((a, eds) <- change.edits if a == name) yield eds).flatten ::: edits)
-      lazy val reverse_edits = edits.reverse
-
-      new Snapshot
-      {
-        val version = stable.current.join
-        val node = version.nodes(name)
-        val is_outdated = !(pending_edits.isEmpty && latest == stable)
-        def lookup_command(id: Command_ID): Option[Command] = state_snapshot.lookup_command(id)
-        def state(command: Command): Command.State =
-          try { state_snapshot.command_state(version, command) }
-          catch { case _: State.Fail => command.empty_state }
-
-        def convert(offset: Text.Offset) = (offset /: edits)((i, edit) => edit.convert(i))
-        def revert(offset: Text.Offset) = (offset /: reverse_edits)((i, edit) => edit.revert(i))
-      }
-    }
-  }
-
-
-
-  /** global state -- document structure and execution process **/
-
   object State
   {
     class Fail(state: State) extends Exception
 
-    val init = State().define_version(Version.init, Map()).assign(Version.init.id, Nil)
+    val init = State().define_version(Version.init, Map()).assign(Version.init.id, Nil)._2
 
     class Assignment(former_assignment: Map[Command, Exec_ID])
     {
@@ -189,6 +188,7 @@
       private val promise = Future.promise[Map[Command, Exec_ID]]
       def is_finished: Boolean = promise.is_finished
       def join: Map[Command, Exec_ID] = promise.join
+      def get_finished: Map[Command, Exec_ID] = promise.get_finished
       def assign(command_execs: List[(Command, Exec_ID)])
       {
         promise.fulfill(tmp_assignment ++ command_execs)
@@ -202,7 +202,8 @@
     val commands: Map[Command_ID, Command.State] = Map(),
     val execs: Map[Exec_ID, (Command.State, Set[Version])] = Map(),
     val assignments: Map[Version, State.Assignment] = Map(),
-    val disposed: Set[ID] = Set())  // FIXME unused!?
+    val disposed: Set[ID] = Set(),  // FIXME unused!?
+    val history: History = History.init)
   {
     private def fail[A]: A = throw new State.Fail(this)
 
@@ -228,7 +229,7 @@
     def the_exec_state(id: Exec_ID): Command.State = execs.getOrElse(id, fail)._1
     def the_assignment(version: Version): State.Assignment = assignments.getOrElse(version, fail)
 
-    def accumulate(id: ID, message: XML.Tree): (Command.State, State) =
+    def accumulate(id: ID, message: XML.Elem): (Command.State, State) =
       execs.get(id) match {
         case Some((st, occs)) =>
           val new_st = st.accumulate(message)
@@ -242,7 +243,7 @@
           }
       }
 
-    def assign(id: Version_ID, edits: List[(Command_ID, Exec_ID)]): State =
+    def assign(id: Version_ID, edits: List[(Command_ID, Exec_ID)]): (List[Command], State) =
     {
       val version = the_version(id)
       val occs = Set(version)  // FIXME unused (!?)
@@ -256,7 +257,7 @@
           (st.command, exec_id)
         }
       the_assignment(version).assign(assigned_execs)  // FIXME explicit value instead of promise (!?)
-      copy(execs = new_execs)
+      (assigned_execs.map(_._1), copy(execs = new_execs))
     }
 
     def is_assigned(version: Version): Boolean =
@@ -265,11 +266,61 @@
         case None => false
       }
 
-    def command_state(version: Version, command: Command): Command.State =
+    def extend_history(previous: Future[Version],
+        edits: List[Node_Text_Edit],
+        result: Future[(List[Edit[Command]], Version)]): (Change, State) =
+    {
+      val change = new Change(previous, edits, result)
+      (change, copy(history = history + change))
+    }
+
+
+    // persistent user-view
+    def snapshot(name: String, pending_edits: List[Text.Edit]): Snapshot =
     {
-      val assgn = the_assignment(version)
-      require(assgn.is_finished)
-      the_exec_state(assgn.join.getOrElse(command, fail))
+      val found_stable = history.undo_list.find(change =>
+        change.is_finished && is_assigned(change.current.get_finished))
+      require(found_stable.isDefined)
+      val stable = found_stable.get
+      val latest = history.undo_list.head
+
+      val edits =
+        (pending_edits /: history.undo_list.takeWhile(_ != stable))((edits, change) =>
+            (for ((a, eds) <- change.edits if a == name) yield eds).flatten ::: edits)
+      lazy val reverse_edits = edits.reverse
+
+      new Snapshot
+      {
+        val version = stable.current.get_finished
+        val node = version.nodes(name)
+        val is_outdated = !(pending_edits.isEmpty && latest == stable)
+
+        def lookup_command(id: Command_ID): Option[Command] = State.this.lookup_command(id)
+
+        def state(command: Command): Command.State =
+          try { the_exec_state(the_assignment(version).get_finished.getOrElse(command, fail)) }
+          catch { case _: State.Fail => command.empty_state }
+
+        def convert(offset: Text.Offset) = (offset /: edits)((i, edit) => edit.convert(i))
+        def revert(offset: Text.Offset) = (offset /: reverse_edits)((i, edit) => edit.revert(i))
+
+        def select_markup[A](range: Text.Range)
+          (result: PartialFunction[Text.Info[Any], A])(default: A): Iterator[Text.Info[A]] =
+        {
+          val former_range = revert(range)
+          for {
+            (command, command_start) <- node.command_range(former_range)
+            Text.Info(r0, x) <- state(command).markup.
+              select((former_range - command_start).restrict(command.range)) {
+                case Text.Info(r0, info)
+                if result.isDefinedAt(Text.Info(convert(r0 + command_start), info)) =>
+                  result(Text.Info(convert(r0 + command_start), info))
+              } { default }
+            val r = convert(r0 + command_start)
+            if !r.is_singularity
+          } yield Text.Info(r, x)
+        }
+      }
     }
   }
 }
--- a/src/Pure/PIDE/isar_document.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/PIDE/isar_document.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -54,6 +54,22 @@
     else if (markup.exists(_.name == Markup.FINISHED)) Finished
     else Unprocessed
   }
+
+
+  /* reported positions */
+
+  def reported_positions(command_id: Document.Command_ID, message: XML.Elem): Set[Text.Range] =
+  {
+    def reported(set: Set[Text.Range], tree: XML.Tree): Set[Text.Range] =
+      tree match {
+        case XML.Elem(Markup(name, Position.Id_Range(id, range)), body)
+        if (name == Markup.BINDING || name == Markup.ENTITY || name == Markup.REPORT) &&
+          id == command_id => body.foldLeft(set + range)(reported)
+        case XML.Elem(_, body) => body.foldLeft(set)(reported)
+        case XML.Text(_) => set
+      }
+    reported(Set.empty, message) ++ Position.Range.unapply(message.markup.properties)
+  }
 }
 
 
--- a/src/Pure/PIDE/markup_tree.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/PIDE/markup_tree.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -90,7 +90,7 @@
     Branches.overlapping(range, branches).toStream
 
   def select[A](root_range: Text.Range)
-    (result: PartialFunction[Text.Info[Any], A])(default: A): Stream[Text.Info[A]] =
+    (result: PartialFunction[Text.Info[Any], A])(default: A): Iterator[Text.Info[A]] =
   {
     def stream(last: Text.Offset, stack: List[(Text.Info[A], Stream[(Text.Range, Branches.Entry)])])
       : Stream[Text.Info[A]] =
@@ -115,10 +115,14 @@
           if (last < stop) parent.restrict(Text.Range(last, stop)) #:: nexts
           else nexts
 
-        case Nil => Stream(Text.Info(Text.Range(last, root_range.stop), default))
+        case Nil =>
+          val stop = root_range.stop
+          if (last < stop) Stream(Text.Info(Text.Range(last, stop), default))
+          else Stream.empty
       }
     }
     stream(root_range.start, List((Text.Info(root_range, default), overlapping(root_range))))
+      .iterator
   }
 
   def swing_tree(parent: DefaultMutableTreeNode)
--- a/src/Pure/PIDE/text.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/PIDE/text.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -33,7 +33,7 @@
     def +(i: Offset): Range = map(_ + i)
     def -(i: Offset): Range = map(_ - i)
 
-    def is_singleton: Boolean = start == stop
+    def is_singularity: Boolean = start == stop
 
     def contains(i: Offset): Boolean = start == i || start < i && i < stop
     def contains(that: Range): Boolean = this.contains(that.start) && that.stop <= this.stop
--- a/src/Pure/Proof/extraction.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Proof/extraction.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -204,7 +204,7 @@
      realizers = Symtab.merge_list (eq_set (op =) o pairself #1) (realizers1, realizers2),
      defs = Library.merge Thm.eq_thm (defs1, defs2),
      expand = Library.merge (op =) (expand1, expand2),
-     prep = (case prep1 of NONE => prep2 | _ => prep1)};
+     prep = if is_some prep1 then prep1 else prep2};
 );
 
 fun read_condeq thy =
--- a/src/Pure/ProofGeneral/proof_general_emacs.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/ProofGeneral/proof_general_emacs.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -10,7 +10,6 @@
   val test_markupN: string
   val sendback: string -> Pretty.T list -> unit
   val init: bool -> unit
-  val init_outer_syntax: unit -> unit
   structure ThyLoad: sig val add_path: string -> unit end
 end;
 
@@ -20,7 +19,6 @@
 
 (* print modes *)
 
-val proof_generalN = "ProofGeneralEmacs";  (*token markup (colouring vars, etc.)*)
 val thm_depsN = "thm_deps";                (*meta-information about theorem deps*)
 val test_markupN = "test_markup";          (*XML markup for everything*)
 
@@ -187,44 +185,35 @@
 
 (* additional outer syntax for Isar *)
 
-fun prP () =
+val _ =
   Outer_Syntax.improper_command "ProofGeneral.pr" "print state (internal)" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o Toplevel.keep (fn state =>
       if Toplevel.is_toplevel state orelse Toplevel.is_theory state then tell_clear_goals ()
       else (Toplevel.quiet := false; Toplevel.print_state true state))));
 
-fun undoP () = (*undo without output -- historical*)
+val _ = (*undo without output -- historical*)
   Outer_Syntax.improper_command "ProofGeneral.undo" "(internal)" Keyword.control
     (Scan.succeed (Toplevel.no_timing o Toplevel.imperative (fn () => Isar.undo 1)));
 
-fun restartP () =
+val _ =
   Outer_Syntax.improper_command "ProofGeneral.restart" "(internal)" Keyword.control
     (Parse.opt_unit >> (Toplevel.no_timing oo K (Toplevel.imperative restart)));
 
-fun kill_proofP () =
+val _ =
   Outer_Syntax.improper_command "ProofGeneral.kill_proof" "(internal)" Keyword.control
     (Scan.succeed (Toplevel.no_timing o
       Toplevel.imperative (fn () => (Isar.kill_proof (); tell_clear_goals ()))));
 
-fun inform_file_processedP () =
+val _ =
   Outer_Syntax.improper_command "ProofGeneral.inform_file_processed" "(internal)" Keyword.control
     (Parse.name >> (fn file =>
       Toplevel.no_timing o Toplevel.imperative (fn () => inform_file_processed file)));
 
-fun inform_file_retractedP () =
+val _ =
   Outer_Syntax.improper_command "ProofGeneral.inform_file_retracted" "(internal)" Keyword.control
     (Parse.name >> (Toplevel.no_timing oo
       (fn file => Toplevel.imperative (fn () => inform_file_retracted file))));
 
-fun process_pgipP () =
-  Outer_Syntax.improper_command "ProofGeneral.process_pgip" "(internal)" Keyword.control
-    (Parse.text >> (Toplevel.no_timing oo
-      (fn txt => Toplevel.imperative (fn () => ProofGeneralPgip.process_pgip txt))));
-
-fun init_outer_syntax () = List.app (fn f => f ())
-  [prP, undoP, restartP, kill_proofP, inform_file_processedP,
-    inform_file_retractedP, process_pgipP];
-
 
 (* init *)
 
@@ -232,18 +221,19 @@
 
 fun init false = panic "No Proof General interface support for Isabelle/classic mode."
   | init true =
-      (! initialized orelse
-        (Output.no_warnings_CRITICAL init_outer_syntax ();
-          Output.add_mode Symbol.xsymbolsN Symbol.output Output.default_escape;
-          Output.add_mode proof_generalN Output.default_output Output.default_escape;
-          Markup.add_mode proof_generalN YXML.output_markup;
-          setup_messages ();
-          ProofGeneralPgip.init_pgip_channel (! Output.priority_fn);
-          setup_thy_loader ();
-          setup_present_hook ();
-          Unsynchronized.set initialized);
-        sync_thy_loader ();
-       Unsynchronized.change print_mode (update (op =) proof_generalN);
+      (if ! initialized then ()
+       else
+        (Output.add_mode Symbol.xsymbolsN Symbol.output Output.default_escape;
+         Output.add_mode ProofGeneralPgip.proof_general_emacsN
+          Output.default_output Output.default_escape;
+         Markup.add_mode ProofGeneralPgip.proof_general_emacsN YXML.output_markup;
+         setup_messages ();
+         ProofGeneralPgip.pgip_channel_emacs (! Output.priority_fn);
+         setup_thy_loader ();
+         setup_present_hook ();
+         initialized := true);
+       sync_thy_loader ();
+       Unsynchronized.change print_mode (update (op =) ProofGeneralPgip.proof_general_emacsN);
        Secure.PG_setup ();
        Isar.toplevel_loop TextIO.stdIn
         {init = true, welcome = true, sync = true, secure = Secure.is_secure ()});
--- a/src/Pure/ProofGeneral/proof_general_keywords.ML	Thu Aug 26 13:15:37 2010 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(*  Title:      Pure/ProofGeneral/proof_general_keywords.ML
-    Author:     Makarius
-
-Dummy session with outer syntax keyword initialization.
-*)
-
-ProofGeneral.init_outer_syntax ();
-
--- a/src/Pure/ProofGeneral/proof_general_pgip.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/ProofGeneral/proof_general_pgip.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -7,12 +7,12 @@
 
 signature PROOF_GENERAL_PGIP =
 sig
+  val proof_general_emacsN: string
+
   val new_thms_deps: theory -> theory -> string list * string list
   val init_pgip: bool -> unit             (* main PGIP loop with true; fail with false *)
 
-  (* These two are just to support the semi-PGIP Emacs mode *)
-  val init_pgip_channel: (string -> unit) -> unit
-  val process_pgip: string -> unit
+  val pgip_channel_emacs: (string -> unit) -> unit
 
   (* More message functions... *)
   val nonfatal_error : string -> unit     (* recoverable (batch) error: carry on scripting *)
@@ -31,6 +31,7 @@
 
 (** print mode **)
 
+val proof_general_emacsN = "ProofGeneralEmacs";
 val proof_generalN = "ProofGeneral";
 val pgmlsymbols_flag = Unsynchronized.ref true;
 
@@ -315,8 +316,6 @@
         fun keyword_elt kind keyword =
             XML.Elem (("keyword", [("word", keyword), ("category", kind)]), [])
         in
-            (* Also, note we don't call init_outer_syntax here to add interface commands,
-            but they should never appear in scripts anyway so it shouldn't matter *)
             Lexicalstructure
               {content = map (keyword_elt "minor") keywords @ map (keyword_elt "major") commands}
         end
@@ -802,7 +801,7 @@
       val filepath = PgipTypes.path_of_pgipurl url
       val filedir = Path.dir filepath
       val thy_name = Path.implode o #1 o Path.split_ext o Path.base
-      val openfile_retract = Output.no_warnings_CRITICAL Thy_Info.kill_thy o thy_name;
+      val openfile_retract = Thy_Info.kill_thy o thy_name;
   in
       case !currently_open_file of
           SOME f => raise PGIP ("<openfile> when a file is already open!\nCurrently open file: " ^
@@ -1007,32 +1006,24 @@
 end
 
 
-(* Extra command for embedding prover-control inside document (obscure/debug usage). *)
-
-fun init_outer_syntax () =
-  Outer_Syntax.improper_command "ProofGeneral.process_pgip" "(internal)" Keyword.control
-    (Parse.text >> (Toplevel.no_timing oo
-      (fn txt => Toplevel.imperative (fn () => process_pgip_plain txt))));
-
-
 (* init *)
 
 val initialized = Unsynchronized.ref false;
 
 fun init_pgip false = panic "No Proof General interface support for Isabelle/classic mode."
   | init_pgip true =
-      (! initialized orelse
+      (if ! initialized then ()
+       else
         (setup_preferences_tweak ();
          Output.add_mode proof_generalN Output.default_output Output.default_escape;
          Markup.add_mode proof_generalN YXML.output_markup;
          setup_messages ();
-         Output.no_warnings_CRITICAL init_outer_syntax ();
          setup_thy_loader ();
          setup_present_hook ();
          init_pgip_session_id ();
          welcome ();
-         Unsynchronized.set initialized);
-        sync_thy_loader ();
+         initialized := true);
+       sync_thy_loader ();
        Unsynchronized.change print_mode (update (op =) proof_generalN);
        pgip_toplevel tty_src);
 
@@ -1045,16 +1036,27 @@
 in
 
 (* Set recipient for PGIP results *)
-fun init_pgip_channel writefn =
+fun pgip_channel_emacs writefn =
     (init_pgip_session_id();
      pgip_output_channel := writefn)
 
 (* Process a PGIP command.
    This works for preferences but not generally guaranteed
    because we haven't done full setup here (e.g., no pgml mode)  *)
-fun process_pgip str =
+fun process_pgip_emacs str =
      setmp_CRITICAL output_xml_fn (!pgip_output_channel) process_pgip_plain str
 
 end
 
+
+(* Extra command for embedding prover-control inside document (obscure/debug usage). *)
+
+val _ =
+  Outer_Syntax.improper_command "ProofGeneral.process_pgip" "(internal)" Keyword.control
+    (Parse.text >> (Toplevel.no_timing oo
+      (fn txt => Toplevel.imperative (fn () =>
+        if print_mode_active proof_general_emacsN
+        then process_pgip_emacs txt
+        else process_pgip_plain txt))));
+
 end;
--- a/src/Pure/ROOT.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/ROOT.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -179,9 +179,9 @@
 use "ML/ml_syntax.ML";
 use "ML/ml_env.ML";
 use "Isar/runtime.ML";
+use "ML/ml_compiler.ML";
 if ml_system = "polyml-5.2" orelse ml_system = "polyml-5.2.1" orelse
-  String.isPrefix "smlnj" ml_system
-then use "ML/ml_compiler.ML"
+  String.isPrefix "smlnj" ml_system then ()
 else use "ML/ml_compiler_polyml-5.3.ML";
 use "ML/ml_context.ML";
 
--- a/src/Pure/Syntax/parser.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Syntax/parser.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -15,7 +15,7 @@
   datatype parsetree =
     Node of string * parsetree list |
     Tip of Lexicon.token
-  val parse: gram -> string -> Lexicon.token list -> parsetree list
+  val parse: Proof.context -> gram -> string -> Lexicon.token list -> parsetree list
   val guess_infix_lr: gram -> string -> (string * bool * bool * int) option
   val branching_level: int Unsynchronized.ref
 end;
@@ -738,7 +738,7 @@
   in get_prods (connected_with chains nts nts) [] end;
 
 
-fun PROCESSS warned prods chains Estate i c states =
+fun PROCESSS ctxt warned prods chains Estate i c states =
   let
     fun all_prods_for nt = prods_for prods chains true c [nt];
 
@@ -772,7 +772,8 @@
                 val dummy =
                   if not (! warned) andalso
                      length (new_states @ States) > ! branching_level then
-                    (warning "Currently parsed expression could be extremely ambiguous.";
+                    (Context_Position.if_visible ctxt warning
+                      "Currently parsed expression could be extremely ambiguous";
                      warned := true)
                   else ();
               in
@@ -809,7 +810,7 @@
   in processS [] states ([], []) end;
 
 
-fun produce warned prods tags chains stateset i indata prev_token =
+fun produce ctxt warned prods tags chains stateset i indata prev_token =
   (case Array.sub (stateset, i) of
     [] =>
       let
@@ -826,16 +827,16 @@
       (case indata of
          [] => Array.sub (stateset, i)
        | c :: cs =>
-         let val (si, sii) = PROCESSS warned prods chains stateset i c s;
+         let val (si, sii) = PROCESSS ctxt warned prods chains stateset i c s;
          in Array.update (stateset, i, si);
             Array.update (stateset, i + 1, sii);
-            produce warned prods tags chains stateset (i + 1) cs c
+            produce ctxt warned prods tags chains stateset (i + 1) cs c
          end));
 
 
 fun get_trees l = map_filter (fn (_, _, [pt], _, _, _) => SOME pt | _ => NONE) l;
 
-fun earley prods tags chains startsymbol indata =
+fun earley ctxt prods tags chains startsymbol indata =
   let
     val start_tag =
       (case Symtab.lookup tags startsymbol of
@@ -846,18 +847,19 @@
     val Estate = Array.array (s, []);
     val _ = Array.update (Estate, 0, S0);
   in
-    get_trees (produce (Unsynchronized.ref false) prods tags chains Estate 0 indata Lexicon.eof)
+    get_trees
+      (produce ctxt (Unsynchronized.ref false) prods tags chains Estate 0 indata Lexicon.eof)
   end;
 
 
-fun parse (Gram {tags, prods, chains, ...}) start toks =
+fun parse ctxt (Gram {tags, prods, chains, ...}) start toks =
   let
     val end_pos =
       (case try List.last toks of
         NONE => Position.none
       | SOME (Lexicon.Token (_, _, (_, end_pos))) => end_pos);
     val r =
-      (case earley prods tags chains start (toks @ [Lexicon.mk_eof end_pos]) of
+      (case earley ctxt prods tags chains start (toks @ [Lexicon.mk_eof end_pos]) of
         [] => raise Fail "Inner syntax: no parse trees"
       | pts => pts);
   in r end;
--- a/src/Pure/Syntax/syntax.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Syntax/syntax.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -485,7 +485,7 @@
     val _ = List.app (Lexicon.report_token ctxt) toks;
 
     val root' = if root <> "prop" andalso is_logtype root then Syn_Ext.logic else root;
-    val pts = Parser.parse gram root' (filter Lexicon.is_proper toks);
+    val pts = Parser.parse ctxt gram root' (filter Lexicon.is_proper toks);
     val len = length pts;
 
     val limit = ! ambiguity_limit;
@@ -495,7 +495,7 @@
     if len <= ! ambiguity_level then ()
     else if ! ambiguity_is_error then error (ambiguity_msg pos)
     else
-      (warning (cat_lines
+      (Context_Position.if_visible ctxt warning (cat_lines
         (("Ambiguous input" ^ Position.str_of pos ^
           "\nproduces " ^ string_of_int len ^ " parse trees" ^
           (if len <= limit then "" else " (" ^ string_of_int limit ^ " displayed)") ^ ":") ::
@@ -519,8 +519,8 @@
 (* read terms *)
 
 (*brute-force disambiguation via type-inference*)
-fun disambig _ _ [t] = t
-  | disambig pp check ts =
+fun disambig _ _ _ [t] = t
+  | disambig ctxt pp check ts =
       let
         val level = ! ambiguity_level;
         val limit = ! ambiguity_limit;
@@ -539,7 +539,8 @@
         if null results then cat_error (ambig_msg ()) (cat_lines (map_filter I errs))
         else if len = 1 then
           (if ambiguity > level then
-            warning "Fortunately, only one parse tree is type correct.\n\
+            Context_Position.if_visible ctxt warning
+              "Fortunately, only one parse tree is type correct.\n\
               \You may still want to disambiguate your grammar or your input."
           else (); hd results)
         else cat_error (ambig_msg ()) (cat_lines
@@ -551,7 +552,7 @@
 fun standard_parse_term pp check get_sort map_const map_free ctxt is_logtype syn ty (syms, pos) =
   read ctxt is_logtype syn ty (syms, pos)
   |> map (Type_Ext.decode_term get_sort map_const map_free)
-  |> disambig (Printer.pp_show_brackets pp) check;
+  |> disambig ctxt (Printer.pp_show_brackets pp) check;
 
 
 (* read types *)
--- a/src/Pure/System/event_bus.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/System/event_bus.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -32,4 +32,29 @@
   /* event invocation */
 
   def event(x: Event) { synchronized { receivers.foreach(_ ! x) } }
+
+
+  /* await global condition -- triggered via bus events */
+
+  def await(cond: => Boolean)
+  {
+    case object Wait
+    val a = new Actor {
+      def act {
+        if (cond) react { case Wait => reply(()); exit(Wait) }
+        else {
+          loop {
+            react {
+              case trigger if trigger != Wait =>
+                if (cond) { react { case Wait => reply(()); exit(Wait) } }
+            }
+          }
+        }
+      }
+    }
+    this += a
+    a.start
+    a !? Wait
+    this -= a
+  }
 }
--- a/src/Pure/System/isabelle_process.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/System/isabelle_process.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -63,7 +63,8 @@
 in
 
 fun standard_message out_stream ch body =
-  message out_stream ch (Position.properties_of (Position.thread_data ())) body;
+  message out_stream ch
+    ((Markup.serialN, serial_string ()) :: Position.properties_of (Position.thread_data ())) body;
 
 fun init_message out_stream =
   message out_stream "A" [(Markup.pidN, process_id ())] (Session.welcome ());
--- a/src/Pure/System/isar.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/System/isar.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -94,7 +94,10 @@
 fun op >> tr =
   (case Toplevel.transition true tr (state ()) of
     NONE => false
-  | SOME (_, SOME err) => (set_exn (SOME err); Toplevel.error_msg tr err; true)
+  | SOME (_, SOME exn_info) =>
+     (set_exn (SOME exn_info);
+      Toplevel.error_msg tr (ML_Compiler.exn_message (Runtime.EXCURSION_FAIL exn_info));
+      true)
   | SOME (st', NONE) =>
       let
         val name = Toplevel.name_of tr;
@@ -134,7 +137,6 @@
 
 fun toplevel_loop in_stream {init = do_init, welcome, sync, secure} =
  (Context.set_thread_data NONE;
-  Secure.Isar_setup ();
   if do_init then init () else ();
   if welcome then writeln (Session.welcome ()) else ();
   uninterruptible (fn _ => fn () => raw_loop secure (Outer_Syntax.isar in_stream sync)) ());
--- a/src/Pure/System/session.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/System/session.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -19,6 +19,7 @@
 
   case object Global_Settings
   case object Perspective
+  case object Assignment
   case class Commands_Changed(set: Set[Command])
 }
 
@@ -44,6 +45,7 @@
   val raw_output = new Event_Bus[Isabelle_Process.Result]
   val commands_changed = new Event_Bus[Session.Commands_Changed]
   val perspective = new Event_Bus[Session.Perspective.type]
+  val assignments = new Event_Bus[Session.Assignment.type]
 
 
   /* unique ids */
@@ -57,17 +59,64 @@
 
 
 
-  /** main actor **/
+  /** buffered command changes (delay_first discipline) **/
+
+  private case object Stop
+
+  private val command_change_buffer = Simple_Thread.actor("command_change_buffer", daemon = true)
+  //{{{
+  {
+    import scala.compat.Platform.currentTime
+
+    var changed: Set[Command] = Set()
+    var flush_time: Option[Long] = None
+
+    def flush_timeout: Long =
+      flush_time match {
+        case None => 5000L
+        case Some(time) => (time - currentTime) max 0
+      }
+
+    def flush()
+    {
+      if (!changed.isEmpty) commands_changed.event(Session.Commands_Changed(changed))
+      changed = Set()
+      flush_time = None
+    }
+
+    def invoke()
+    {
+      val now = currentTime
+      flush_time match {
+        case None => flush_time = Some(now + output_delay)
+        case Some(time) => if (now >= time) flush()
+      }
+    }
+
+    var finished = false
+    while (!finished) {
+      receiveWithin(flush_timeout) {
+        case command: Command => changed += command; invoke()
+        case TIMEOUT => flush()
+        case Stop => finished = true
+        case bad => System.err.println("command_change_buffer: ignoring bad message " + bad)
+      }
+    }
+  }
+  //}}}
+
+
+
+  /** main protocol actor **/
 
   @volatile private var syntax = new Outer_Syntax(system.symbols)
   def current_syntax(): Outer_Syntax = syntax
 
-  @volatile private var global_state = Document.State.init
-  private def change_state(f: Document.State => Document.State) { global_state = f(global_state) }
-  def current_state(): Document.State = global_state
+  private val global_state = new Volatile(Document.State.init)
+  def current_state(): Document.State = global_state.peek()
 
+  private case class Edit_Version(edits: List[Document.Node_Text_Edit])
   private case class Started(timeout: Int, args: List[String])
-  private case object Stop
 
   private val session_actor = Simple_Thread.actor("session_actor", daemon = true)
   {
@@ -79,12 +128,10 @@
     def handle_change(change: Document.Change)
     //{{{
     {
-      require(change.is_finished)
+      val previous = change.previous.get_finished
+      val (node_edits, current) = change.result.get_finished
 
-      val previous = change.previous.join
-      val (node_edits, current) = change.result.join
-
-      var former_assignment = current_state().the_assignment(previous).join
+      var former_assignment = global_state.peek().the_assignment(previous).get_finished
       for {
         (name, Some(cmd_edits)) <- node_edits
         (prev, None) <- cmd_edits
@@ -103,8 +150,8 @@
                     c2 match {
                       case None => None
                       case Some(command) =>
-                        if (current_state().lookup_command(command.id).isEmpty) {
-                          change_state(_.define_command(command))
+                        if (global_state.peek().lookup_command(command.id).isEmpty) {
+                          global_state.change(_.define_command(command))
                           prover.define_command(command.id, system.symbols.encode(command.source))
                         }
                         Some(command.id)
@@ -113,7 +160,7 @@
               }
             (name -> Some(ids))
         }
-      change_state(_.define_version(current, former_assignment))
+      global_state.change(_.define_version(current, former_assignment))
       prover.edit_version(previous.id, current.id, id_edits)
     }
     //}}}
@@ -131,19 +178,22 @@
     {
       raw_protocol.event(result)
 
-      Position.get_id(result.properties) match {
-        case Some(state_id) =>
+      result.properties match {
+        case Position.Id(state_id) =>
           try {
-            val (st, state) = global_state.accumulate(state_id, result.message)
-            global_state = state
-            indicate_command_change(st.command)
+            val st = global_state.change_yield(_.accumulate(state_id, result.message))
+            command_change_buffer ! st.command
           }
           catch { case _: Document.State.Fail => bad_result(result) }
-        case None =>
+        case _ =>
           if (result.is_status) {
             result.body match {
               case List(Isar_Document.Assign(id, edits)) =>
-                try { change_state(_.assign(id, edits)) }
+                try {
+                  val cmds: List[Command] = global_state.change_yield(_.assign(id, edits))
+                  for (cmd <- cmds) command_change_buffer ! cmd
+                  assignments.event(Session.Assignment)
+                }
                 catch { case _: Document.State.Fail => bad_result(result) }
               case List(Keyword.Command_Decl(name, kind)) => syntax += (name, kind)
               case List(Keyword.Keyword_Decl(name)) => syntax += name
@@ -202,6 +252,24 @@
     var finished = false
     while (!finished) {
       receive {
+        case Edit_Version(edits) =>
+          val previous = global_state.peek().history.tip.current
+          val result = Future.fork { Thy_Syntax.text_edits(Session.this, previous.join, edits) }
+          val change = global_state.change_yield(_.extend_history(previous, edits, result))
+
+          val this_actor = self
+          change.current.map(_ => {
+            assignments.await { global_state.peek().is_assigned(previous.get_finished) }
+            this_actor ! change })
+
+          reply(())
+
+        case change: Document.Change if prover != null =>
+          handle_change(change)
+
+        case result: Isabelle_Process.Result =>
+          handle_result(result)
+
         case Started(timeout, args) =>
           if (prover == null) {
             prover = new Isabelle_Process(system, self, args:_*) with Isar_Document
@@ -219,13 +287,7 @@
             finished = true
           }
 
-        case change: Document.Change if prover != null =>
-          handle_change(change)
-
-        case result: Isabelle_Process.Result =>
-          handle_result(result)
-
-        case TIMEOUT =>  // FIXME clarify!
+        case TIMEOUT =>  // FIXME clarify
 
         case bad if prover != null =>
           System.err.println("session_actor: ignoring bad message " + bad)
@@ -235,95 +297,15 @@
 
 
 
-  /** buffered command changes (delay_first discipline) **/
-
-  private val command_change_buffer = actor
-  //{{{
-  {
-    import scala.compat.Platform.currentTime
-
-    var changed: Set[Command] = Set()
-    var flush_time: Option[Long] = None
-
-    def flush_timeout: Long =
-      flush_time match {
-        case None => 5000L
-        case Some(time) => (time - currentTime) max 0
-      }
-
-    def flush()
-    {
-      if (!changed.isEmpty) commands_changed.event(Session.Commands_Changed(changed))
-      changed = Set()
-      flush_time = None
-    }
-
-    def invoke()
-    {
-      val now = currentTime
-      flush_time match {
-        case None => flush_time = Some(now + output_delay)
-        case Some(time) => if (now >= time) flush()
-      }
-    }
-
-    loop {
-      reactWithin(flush_timeout) {
-        case command: Command => changed += command; invoke()
-        case TIMEOUT => flush()
-        case bad => System.err.println("command_change_buffer: ignoring bad message " + bad)
-      }
-    }
-  }
-  //}}}
-
-  def indicate_command_change(command: Command)
-  {
-    command_change_buffer ! command
-  }
-
-
-
-  /** editor history **/
-
-  private case class Edit_Version(edits: List[Document.Node_Text_Edit])
-
-  @volatile private var history = Document.History.init
-
-  def snapshot(name: String, pending_edits: List[Text.Edit]): Document.Snapshot =
-    history.snapshot(name, pending_edits, current_state())
-
-  private val editor_history = actor
-  {
-    loop {
-      react {
-        case Edit_Version(edits) =>
-          val prev = history.tip.current
-          val result =
-            // FIXME potential denial-of-service concerning worker pool (!?!?)
-            isabelle.Future.fork {
-              val previous = prev.join
-              val former_assignment = current_state().the_assignment(previous).join  // FIXME async!?
-              Thy_Syntax.text_edits(Session.this, previous, edits)
-            }
-          val change = new Document.Change(prev, edits, result)
-          history += change
-          change.current.map(_ => session_actor ! change)
-          reply(())
-
-        case bad => System.err.println("editor_model: ignoring bad message " + bad)
-      }
-    }
-  }
-
-
-
   /** main methods **/
 
   def started(timeout: Int, args: List[String]): Option[String] =
     (session_actor !? Started(timeout, args)).asInstanceOf[Option[String]]
 
-  def stop() { session_actor ! Stop }
+  def stop() { command_change_buffer ! Stop; session_actor ! Stop }
+
+  def edit_version(edits: List[Document.Node_Text_Edit]) { session_actor !? Edit_Version(edits) }
 
-  def edit_version(edits: List[Document.Node_Text_Edit]) { editor_history !? Edit_Version(edits) }
+  def snapshot(name: String, pending_edits: List[Text.Edit]): Document.Snapshot =
+    global_state.peek().snapshot(name, pending_edits)
 }
--- a/src/Pure/System/swing_thread.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/System/swing_thread.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -23,7 +23,7 @@
 
   def now[A](body: => A): A =
   {
-    var result: Option[A] = None
+    @volatile var result: Option[A] = None
     if (SwingUtilities.isEventDispatchThread()) { result = Some(body) }
     else SwingUtilities.invokeAndWait(new Runnable { def run = { result = Some(body) } })
     result.get
--- a/src/Pure/Thy/thy_load.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Thy/thy_load.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -195,7 +195,7 @@
       val _ = Context.>> Local_Theory.propagate_ml_env;
 
       val provide = provide (src_path, (path, id));
-      val _ = Context.>> (Context.mapping provide (Local_Theory.theory provide));
+      val _ = Context.>> (Context.mapping provide (Local_Theory.background_theory provide));
     in () end;
 
 fun exec_ml src_path = ML_Context.exec (fn () => use_ml src_path);
--- a/src/Pure/Thy/thy_output.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Thy/thy_output.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -6,31 +6,37 @@
 
 signature THY_OUTPUT =
 sig
-  val display: bool Unsynchronized.ref
-  val quotes: bool Unsynchronized.ref
-  val indent: int Unsynchronized.ref
-  val source: bool Unsynchronized.ref
-  val break: bool Unsynchronized.ref
-  val add_commands: (string * (Args.src -> Toplevel.state -> string)) list -> unit
-  val add_options: (string * (string -> (unit -> string) -> unit -> string)) list -> unit
+  val display_default: bool Unsynchronized.ref
+  val quotes_default: bool Unsynchronized.ref
+  val indent_default: int Unsynchronized.ref
+  val source_default: bool Unsynchronized.ref
+  val break_default: bool Unsynchronized.ref
+  val display: bool Config.T
+  val quotes: bool Config.T
+  val indent: int Config.T
+  val source: bool Config.T
+  val break: bool Config.T
+  val add_wrapper: ((unit -> string) -> unit -> string) -> Proof.context -> Proof.context
+  val add_option: string -> (string -> Proof.context -> Proof.context) -> unit
   val defined_command: string -> bool
   val defined_option: string -> bool
   val print_antiquotations: unit -> unit
   val boolean: string -> bool
   val integer: string -> int
   val antiquotation: string -> 'a context_parser ->
-    ({state: Toplevel.state, source: Args.src, context: Proof.context} -> 'a -> string) -> unit
+    ({source: Args.src, state: Toplevel.state, context: Proof.context} -> 'a -> string) -> unit
   datatype markup = Markup | MarkupEnv | Verbatim
   val modes: string list Unsynchronized.ref
   val eval_antiquote: Scan.lexicon -> Toplevel.state -> Symbol_Pos.text * Position.T -> string
   val present_thy: Scan.lexicon -> (string -> string list) -> (markup -> string -> bool) ->
     (Toplevel.transition * Toplevel.state) list -> (Token.T, 'a) Source.source -> Buffer.T
-  val pretty_text: string -> Pretty.T
+  val pretty_text: Proof.context -> string -> Pretty.T
   val pretty_term: Proof.context -> term -> Pretty.T
   val pretty_thm: Proof.context -> thm -> Pretty.T
   val str_of_source: Args.src -> string
-  val maybe_pretty_source: ('a -> Pretty.T) -> Args.src -> 'a list -> Pretty.T list
-  val output: Pretty.T list -> string
+  val maybe_pretty_source: (Proof.context -> 'a -> Pretty.T) -> Proof.context ->
+    Args.src -> 'a list -> Pretty.T list
+  val output: Proof.context -> Pretty.T list -> string
 end;
 
 structure Thy_Output: THY_OUTPUT =
@@ -38,11 +44,31 @@
 
 (** global options **)
 
-val display = Unsynchronized.ref false;
-val quotes = Unsynchronized.ref false;
-val indent = Unsynchronized.ref 0;
-val source = Unsynchronized.ref false;
-val break = Unsynchronized.ref false;
+val display_default = Unsynchronized.ref false;
+val quotes_default = Unsynchronized.ref false;
+val indent_default = Unsynchronized.ref 0;
+val source_default = Unsynchronized.ref false;
+val break_default = Unsynchronized.ref false;
+
+val (display, setup_display) = Attrib.config_bool "thy_output_display" (fn _ => ! display_default);
+val (quotes, setup_quotes) = Attrib.config_bool "thy_output_quotes" (fn _ => ! quotes_default);
+val (indent, setup_indent) = Attrib.config_int "thy_output_indent" (fn _ => ! indent_default);
+val (source, setup_source) = Attrib.config_bool "thy_output_source" (fn _ => ! source_default);
+val (break, setup_break) = Attrib.config_bool "thy_output_break" (fn _ => ! break_default);
+
+val _ = Context.>> (Context.map_theory
+ (setup_display #> setup_quotes #> setup_indent #> setup_source #> setup_break));
+
+
+structure Wrappers = Proof_Data
+(
+  type T = ((unit -> string) -> unit -> string) list;
+  fun init _ = [];
+);
+
+fun add_wrapper wrapper = Wrappers.map (cons wrapper);
+
+val wrap = Wrappers.get #> fold (fn wrapper => fn f => wrapper f);
 
 
 
@@ -51,22 +77,23 @@
 local
 
 val global_commands =
-  Unsynchronized.ref (Symtab.empty: (Args.src -> Toplevel.state -> string) Symtab.table);
+  Unsynchronized.ref
+    (Symtab.empty: (Args.src -> Toplevel.state -> Proof.context -> string) Symtab.table);
 
 val global_options =
-  Unsynchronized.ref (Symtab.empty: (string -> (unit -> string) -> unit -> string) Symtab.table);
+  Unsynchronized.ref (Symtab.empty: (string -> Proof.context -> Proof.context) Symtab.table);
 
-fun add_item kind (name, x) tab =
+fun add_item kind name item tab =
  (if not (Symtab.defined tab name) then ()
   else warning ("Redefined document antiquotation " ^ kind ^ ": " ^ quote name);
-  Symtab.update (name, x) tab);
+  Symtab.update (name, item) tab);
 
 in
 
-fun add_commands xs =
-  CRITICAL (fn () => Unsynchronized.change global_commands (fold (add_item "command") xs));
-fun add_options xs =
-  CRITICAL (fn () => Unsynchronized.change global_options (fold (add_item "option") xs));
+fun add_command name cmd =
+  CRITICAL (fn () => Unsynchronized.change global_commands (add_item "command" name cmd));
+fun add_option name opt =
+  CRITICAL (fn () => Unsynchronized.change global_options (add_item "option" name opt));
 
 fun defined_command name = Symtab.defined (! global_commands) name;
 fun defined_option name = Symtab.defined (! global_options) name;
@@ -77,18 +104,15 @@
       NONE => error ("Unknown document antiquotation: " ^ quote name ^ Position.str_of pos)
     | SOME f =>
        (Position.report (Markup.doc_antiq name) pos;
-        (fn state => f src state handle ERROR msg =>
+        (fn state => fn ctxt => f src state ctxt handle ERROR msg =>
           cat_error msg ("The error(s) above occurred in document antiquotation: " ^
             quote name ^ Position.str_of pos))))
   end;
 
-fun option (name, s) f () =
+fun option (name, s) ctxt =
   (case Symtab.lookup (! global_options) name of
     NONE => error ("Unknown document antiquotation option: " ^ quote name)
-  | SOME opt => opt s f ());
-
-fun options [] f = f
-  | options (opt :: opts) f = option opt (options opts f);
+  | SOME opt => opt s ctxt);
 
 
 fun print_antiquotations () =
@@ -100,10 +124,11 @@
 
 end;
 
-fun antiquotation name scan out = add_commands [(name, fn src => fn state =>
-  let val (x, ctxt) = Args.context_syntax "document antiquotation"
-    scan src (Toplevel.presentation_context_of state)
-  in out {source = src, state = state, context = ctxt} x end)];
+fun antiquotation name scan out =
+  add_command name
+    (fn src => fn state => fn ctxt =>
+      let val (x, ctxt') = Args.context_syntax "document antiquotation" scan src ctxt
+      in out {source = src, state = state, context = ctxt'} x end);
 
 
 
@@ -151,11 +176,13 @@
   let
     fun expand (Antiquote.Text ss) = Symbol_Pos.content ss
       | expand (Antiquote.Antiq (ss, (pos, _))) =
-          let val (opts, src) = Token.read_antiq lex antiq (ss, pos) in
-            options opts (fn () => command src state) ();  (*preview errors!*)
-            Print_Mode.with_modes (! modes @ Latex.modes)
-              (Output.no_warnings_CRITICAL (options opts (fn () => command src state))) ()
-          end
+          let
+            val (opts, src) = Token.read_antiq lex antiq (ss, pos);
+            fun cmd ctxt = wrap ctxt (fn () => command src state ctxt) ();
+            val preview_ctxt = fold option opts (Toplevel.presentation_context_of state);
+            val print_ctxt = Context_Position.set_visible false preview_ctxt;
+            val _ = cmd preview_ctxt;
+          in Print_Mode.with_modes (! modes @ Latex.modes) (fn () => cmd print_ctxt) () end
       | expand (Antiquote.Open _) = ""
       | expand (Antiquote.Close _) = "";
     val ants = Antiquote.read (Symbol_Pos.explode (txt, pos), pos);
@@ -417,31 +444,31 @@
 
 (* options *)
 
-val _ = add_options
- [("show_types", setmp_CRITICAL Syntax.show_types o boolean),
-  ("show_sorts", setmp_CRITICAL Syntax.show_sorts o boolean),
-  ("show_structs", setmp_CRITICAL show_structs o boolean),
-  ("show_question_marks", setmp_CRITICAL show_question_marks o boolean),
-  ("long_names", setmp_CRITICAL Name_Space.long_names o boolean),
-  ("short_names", setmp_CRITICAL Name_Space.short_names o boolean),
-  ("unique_names", setmp_CRITICAL Name_Space.unique_names o boolean),
-  ("eta_contract", setmp_CRITICAL Syntax.eta_contract o boolean),
-  ("display", setmp_CRITICAL display o boolean),
-  ("break", setmp_CRITICAL break o boolean),
-  ("quotes", setmp_CRITICAL quotes o boolean),
-  ("mode", fn s => fn f => Print_Mode.with_modes [s] f),
-  ("margin", setmp_CRITICAL Pretty.margin_default o integer),
-  ("indent", setmp_CRITICAL indent o integer),
-  ("source", setmp_CRITICAL source o boolean),
-  ("goals_limit", setmp_CRITICAL goals_limit o integer)];
+val _ = add_option "show_types" (add_wrapper o setmp_CRITICAL Syntax.show_types o boolean);
+val _ = add_option "show_sorts" (add_wrapper o setmp_CRITICAL Syntax.show_sorts o boolean);
+val _ = add_option "show_structs" (add_wrapper o setmp_CRITICAL show_structs o boolean);
+val _ = add_option "show_question_marks" (add_wrapper o setmp_CRITICAL show_question_marks o boolean);
+val _ = add_option "long_names" (add_wrapper o setmp_CRITICAL Name_Space.long_names o boolean);
+val _ = add_option "short_names" (add_wrapper o setmp_CRITICAL Name_Space.short_names o boolean);
+val _ = add_option "unique_names" (add_wrapper o setmp_CRITICAL Name_Space.unique_names o boolean);
+val _ = add_option "eta_contract" (add_wrapper o setmp_CRITICAL Syntax.eta_contract o boolean);
+val _ = add_option "display" (Config.put display o boolean);
+val _ = add_option "break" (Config.put break o boolean);
+val _ = add_option "quotes" (Config.put quotes o boolean);
+val _ = add_option "mode" (add_wrapper o Print_Mode.with_modes o single);
+val _ = add_option "margin" (add_wrapper o setmp_CRITICAL Pretty.margin_default o integer);
+val _ = add_option "indent" (Config.put indent o integer);
+val _ = add_option "source" (Config.put source o boolean);
+val _ = add_option "goals_limit" (add_wrapper o setmp_CRITICAL goals_limit o integer);
 
 
 (* basic pretty printing *)
 
-fun tweak_line s =
-  if ! display then s else Symbol.strip_blanks s;
+fun tweak_line ctxt s =
+  if Config.get ctxt display then s else Symbol.strip_blanks s;
 
-val pretty_text = Pretty.chunks o map Pretty.str o map tweak_line o Library.split_lines;
+fun pretty_text ctxt =
+  Pretty.chunks o map Pretty.str o map (tweak_line ctxt) o Library.split_lines;
 
 fun pretty_term ctxt t = Syntax.pretty_term (Variable.auto_fixes t ctxt) t;
 
@@ -490,19 +517,19 @@
 
 val str_of_source = space_implode " " o map Token.unparse o #2 o #1 o Args.dest_src;
 
-fun maybe_pretty_source pretty src xs =
-  map pretty xs  (*always pretty in order to exhibit errors!*)
-  |> (if ! source then K [pretty_text (str_of_source src)] else I);
+fun maybe_pretty_source pretty ctxt src xs =
+  map (pretty ctxt) xs  (*always pretty in order to exhibit errors!*)
+  |> (if Config.get ctxt source then K [pretty_text ctxt (str_of_source src)] else I);
 
-fun output prts =
+fun output ctxt prts =
   prts
-  |> (if ! quotes then map Pretty.quote else I)
-  |> (if ! display then
-    map (Output.output o Pretty.string_of o Pretty.indent (! indent))
+  |> (if Config.get ctxt quotes then map Pretty.quote else I)
+  |> (if Config.get ctxt display then
+    map (Output.output o Pretty.string_of o Pretty.indent (Config.get ctxt indent))
     #> space_implode "\\isasep\\isanewline%\n"
     #> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
   else
-    map (Output.output o (if ! break then Pretty.string_of else Pretty.str_of))
+    map (Output.output o (if Config.get ctxt break then Pretty.string_of else Pretty.str_of))
     #> space_implode "\\isasep\\isanewline%\n"
     #> enclose "\\isa{" "}");
 
@@ -515,11 +542,12 @@
 local
 
 fun basic_entities name scan pretty = antiquotation name scan
-  (fn {source, context, ...} => output o maybe_pretty_source (pretty context) source);
+  (fn {source, context, ...} => output context o maybe_pretty_source pretty context source);
 
 fun basic_entities_style name scan pretty = antiquotation name scan
   (fn {source, context, ...} => fn (style, xs) =>
-    output (maybe_pretty_source (fn x => pretty context (style, x)) source xs));
+    output context
+      (maybe_pretty_source (fn ctxt => fn x => pretty ctxt (style, x)) context source xs));
 
 fun basic_entity name scan = basic_entities name (scan >> single);
 
@@ -533,7 +561,7 @@
 val _ = basic_entity "const" (Args.const_proper false) pretty_const;
 val _ = basic_entity "abbrev" (Scan.lift Args.name_source) pretty_abbrev;
 val _ = basic_entity "typ" Args.typ_abbrev Syntax.pretty_typ;
-val _ = basic_entity "text" (Scan.lift Args.name) (K pretty_text);
+val _ = basic_entity "text" (Scan.lift Args.name) pretty_text;
 val _ = basic_entities "prf" Attrib.thms (pretty_prf false);
 val _ = basic_entities "full_prf" Attrib.thms (pretty_prf true);
 val _ = basic_entity "theory" (Scan.lift Args.name) pretty_theory;
@@ -554,7 +582,7 @@
   | _ => error "No proof state");
 
 fun goal_state name main_goal = antiquotation name (Scan.succeed ())
-  (fn {state, ...} => fn () => output
+  (fn {state, context, ...} => fn () => output context
     [Pretty.chunks (Proof.pretty_goals main_goal (proof_state state))]);
 
 in
@@ -578,7 +606,7 @@
       val _ = context
         |> Proof.theorem NONE (K I) [[(prop, [])]]
         |> Proof.global_terminal_proof methods;
-    in output (maybe_pretty_source (pretty_term context) prop_src [prop]) end);
+    in output context (maybe_pretty_source pretty_term context prop_src [prop]) end);
 
 
 (* ML text *)
@@ -589,8 +617,8 @@
   (fn {context, ...} => fn (txt, pos) =>
    (ML_Context.eval_in (SOME context) false pos (ml pos txt);
     Symbol_Pos.content (Symbol_Pos.explode (txt, pos))
-    |> (if ! quotes then quote else I)
-    |> (if ! display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"
+    |> (if Config.get context quotes then quote else I)
+    |> (if Config.get context display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"
     else
       split_lines
       #> map (space_implode "\\verb,|," o map (enclose "\\verb|" "|") o space_explode "|")
--- a/src/Pure/Thy/thy_syntax.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/Thy/thy_syntax.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -77,9 +77,11 @@
       commands.iterator.find(_.is_unparsed) match {
         case Some(first_unparsed) =>
           val first =
-            commands.reverse_iterator(first_unparsed).find(_.is_command) getOrElse commands.head
+            commands.reverse_iterator(first_unparsed).
+              dropWhile(_.newlines == 0).find(_.is_command) getOrElse commands.head
           val last =
-            commands.iterator(first_unparsed).find(_.is_command) getOrElse commands.last
+            commands.iterator(first_unparsed).
+              dropWhile(_.newlines == 0).find(_.is_command) getOrElse commands.last
           val range =
             commands.iterator(first).takeWhile(_ != last).toList ::: List(last)
 
--- a/src/Pure/build-jars	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/build-jars	Wed Sep 01 17:19:47 2010 +0200
@@ -24,6 +24,7 @@
 declare -a SOURCES=(
   Concurrent/future.scala
   Concurrent/simple_thread.scala
+  Concurrent/volatile.scala
   General/exn.scala
   General/linear_set.scala
   General/markup.scala
--- a/src/Pure/config.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/config.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -41,8 +41,8 @@
   | print_value (Int i) = signed_string_of_int i
   | print_value (String s) = quote s;
 
-fun print_type (Bool _) = "boolean"
-  | print_type (Int _) = "integer"
+fun print_type (Bool _) = "bool"
+  | print_type (Int _) = "int"
   | print_type (String _) = "string";
 
 fun same_type (Bool _) (Bool _) = true
--- a/src/Pure/context_position.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/context_position.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -9,6 +9,7 @@
   val is_visible: Proof.context -> bool
   val set_visible: bool -> Proof.context -> Proof.context
   val restore_visible: Proof.context -> Proof.context -> Proof.context
+  val if_visible: Proof.context -> ('a -> unit) -> 'a -> unit
   val report_text: Proof.context -> Markup.T -> Position.T -> string -> unit
   val report: Proof.context -> Markup.T -> Position.T -> unit
 end;
@@ -26,6 +27,8 @@
 val set_visible = Data.put;
 val restore_visible = set_visible o is_visible;
 
+fun if_visible ctxt f x = if is_visible ctxt then f x else ();
+
 fun report_text ctxt markup pos txt =
   if is_visible ctxt then Position.report_text markup pos txt else ();
 
--- a/src/Pure/goal.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/goal.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -159,7 +159,7 @@
   (case SINGLE (tac (map Assumption.assume casms)) (init cprop) of
     SOME th => Drule.implies_intr_list casms
       (finish (Syntax.init_pretty_global (Thm.theory_of_thm th)) th)
-  | NONE => error "Tactic failed.");
+  | NONE => error "Tactic failed");
 
 
 (* prove_common etc. *)
@@ -191,7 +191,7 @@
 
     fun result () =
       (case SINGLE (tac {prems = prems, context = ctxt'}) (init stmt) of
-        NONE => err "Tactic failed."
+        NONE => err "Tactic failed"
       | SOME st =>
           let val res = finish ctxt' st handle THM (msg, _, _) => err msg in
             if Unify.matches_list thy [Thm.term_of stmt] [Thm.prop_of res]
--- a/src/Pure/meta_simplifier.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/meta_simplifier.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -278,9 +278,19 @@
 
 val trace_simp_default = Unsynchronized.ref false;
 val trace_simp_value =
-  Config.declare false "trace_simp" (fn _ => Config.Bool (!trace_simp_default));
+  Config.declare false "trace_simp" (fn _ => Config.Bool (! trace_simp_default));
 val trace_simp = Config.bool trace_simp_value;
 
+fun if_enabled (Simpset ({context, ...}, _)) flag f =
+  (case context of
+    SOME ctxt => if Config.get ctxt flag then f ctxt else ()
+  | NONE => ())
+
+fun if_visible (Simpset ({context, ...}, _)) f x =
+  (case context of
+    SOME ctxt => if Context_Position.is_visible ctxt then f x else ()
+  | NONE => ());
+
 local
 
 fun prnt ss warn a = if warn then warning a else trace_depth ss a;
@@ -301,11 +311,6 @@
 fun print_term_global ss warn a thy t =
   print_term ss warn (K a) t (ProofContext.init_global thy);
 
-fun if_enabled (Simpset ({context, ...}, _)) flag f =
-  (case context of
-    SOME ctxt => if Config.get ctxt flag then f ctxt else ()
-  | NONE => ())
-
 fun debug warn a ss = if_enabled ss debug_simp (fn _ => prnt ss warn (a ()));
 fun trace warn a ss = if_enabled ss trace_simp (fn _ => prnt ss warn (a ()));
 
@@ -326,9 +331,7 @@
 fun warn_thm a ss th =
   print_term_global ss true a (Thm.theory_of_thm th) (Thm.full_prop_of th);
 
-fun cond_warn_thm a (ss as Simpset ({context, ...}, _)) th =
-  if (case context of NONE => true | SOME ctxt => Context_Position.is_visible ctxt)
-  then warn_thm a ss th else ();
+fun cond_warn_thm a ss th = if_visible ss (fn () => warn_thm a ss th) ();
 
 end;
 
@@ -571,8 +574,9 @@
       val a = the (cong_name (head_of (term_of lhs))) handle Option.Option =>
         raise SIMPLIFIER ("Congruence must start with a constant or free variable", thm);
       val (xs, weak) = congs;
-      val _ =  if AList.defined (op =) xs a
-        then warning ("Overwriting congruence rule for " ^ quote a)
+      val _ =
+        if AList.defined (op =) xs a
+        then if_visible ss warning ("Overwriting congruence rule for " ^ quote a)
         else ();
       val xs' = AList.update (op =) (a, thm) xs;
       val weak' = if is_full_cong thm then weak else a :: weak;
@@ -643,14 +647,14 @@
     (congs, Net.insert_term eq_proc (term_of lhs, proc) procs,
       mk_rews, termless, subgoal_tac, loop_tacs, solvers)) ss
   handle Net.INSERT =>
-    (warning ("Ignoring duplicate simplification procedure " ^ quote name); ss));
+    (if_visible ss warning ("Ignoring duplicate simplification procedure " ^ quote name); ss));
 
 fun del_proc (proc as Proc {name, lhs, ...}) ss =
   map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
     (congs, Net.delete_term eq_proc (term_of lhs, proc) procs,
       mk_rews, termless, subgoal_tac, loop_tacs, solvers)) ss
   handle Net.DELETE =>
-    (warning ("Simplification procedure " ^ quote name ^ " not in simpset"); ss);
+    (if_visible ss warning ("Simplification procedure " ^ quote name ^ " not in simpset"); ss);
 
 fun prep_procs (Simproc {name, lhss, proc, id}) =
   lhss |> map (fn lhs => Proc {name = name, lhs = lhs, proc = Morphism.form proc, id = id});
@@ -720,20 +724,18 @@
 fun ss addloop' (name, tac) = ss |>
   map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
     (congs, procs, mk_rews, termless, subgoal_tac,
-      (if AList.defined (op =) loop_tacs name
-        then warning ("Overwriting looper " ^ quote name)
-        else (); AList.update (op =) (name, tac) loop_tacs),
-      solvers));
+     (if AList.defined (op =) loop_tacs name
+      then if_visible ss warning ("Overwriting looper " ^ quote name)
+      else (); AList.update (op =) (name, tac) loop_tacs), solvers));
 
 fun ss addloop (name, tac) = ss addloop' (name, K tac);
 
 fun ss delloop name = ss |>
   map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
     (congs, procs, mk_rews, termless, subgoal_tac,
-      (if AList.defined (op =) loop_tacs name
-        then ()
-        else warning ("No such looper in simpset: " ^ quote name);
-       AList.delete (op =) name loop_tacs), solvers));
+     (if AList.defined (op =) loop_tacs name then ()
+      else if_visible ss warning ("No such looper in simpset: " ^ quote name);
+      AList.delete (op =) name loop_tacs), solvers));
 
 fun ss setSSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless,
   subgoal_tac, loop_tacs, (unsafe_solvers, _)) =>
@@ -851,7 +853,7 @@
 fun mk_procrule ss thm =
   let val (_, prems, lhs, elhs, rhs, _) = decomp_simp thm in
     if rewrite_rule_extra_vars prems lhs rhs
-    then (warn_thm "Extra vars on rhs:" ss thm; [])
+    then (cond_warn_thm "Extra vars on rhs:" ss thm; [])
     else [mk_rrule2 {thm = thm, name = "", lhs = lhs, elhs = elhs, perm = false}]
   end;
 
@@ -922,14 +924,19 @@
            if unconditional
            then
              (trace_thm (fn () => "Rewriting:") ss thm';
-              let val lr = Logic.dest_equals prop;
-                  val SOME thm'' = check_conv false ss eta_thm thm'
+              let
+                val lr = Logic.dest_equals prop;
+                val SOME thm'' = check_conv false ss eta_thm thm';
               in SOME (thm'', uncond_skel (congs, lr)) end)
            else
              (trace_thm (fn () => "Trying to rewrite:") ss thm';
               if simp_depth ss > Config.get ctxt simp_depth_limit
-              then let val s = "simp_depth_limit exceeded - giving up"
-                   in trace false (fn () => s) ss; warning s; NONE end
+              then
+                let
+                  val s = "simp_depth_limit exceeded - giving up";
+                  val _ = trace false (fn () => s) ss;
+                  val _ = if_visible ss warning s;
+                in NONE end
               else
               case prover ss thm' of
                 NONE => (trace_thm (fn () => "FAILED") ss thm'; NONE)
@@ -947,15 +954,18 @@
           let val opt = rew rrule handle Pattern.MATCH => NONE
           in case opt of NONE => rews rrules | some => some end;
 
-    fun sort_rrules rrs = let
-      fun is_simple({thm, ...}:rrule) = case Thm.prop_of thm of
-                                      Const("==",_) $ _ $ _ => true
-                                      | _                   => false
-      fun sort []        (re1,re2) = re1 @ re2
-        | sort (rr::rrs) (re1,re2) = if is_simple rr
-                                     then sort rrs (rr::re1,re2)
-                                     else sort rrs (re1,rr::re2)
-    in sort rrs ([],[]) end
+    fun sort_rrules rrs =
+      let
+        fun is_simple ({thm, ...}: rrule) =
+          (case Thm.prop_of thm of
+            Const ("==", _) $ _ $ _ => true
+          | _ => false);
+        fun sort [] (re1, re2) = re1 @ re2
+          | sort (rr :: rrs) (re1, re2) =
+              if is_simple rr
+              then sort rrs (rr :: re1, re2)
+              else sort rrs (re1, rr :: re2);
+      in sort rrs ([], []) end;
 
     fun proc_rews [] = NONE
       | proc_rews (Proc {name, proc, lhs, ...} :: ps) =
@@ -971,12 +981,13 @@
                       " -- does not match") ss t; proc_rews ps)
                   | some => some)))
           else proc_rews ps;
-  in case eta_t of
-       Abs _ $ _ => SOME (Thm.transitive eta_thm
-         (Thm.beta_conversion false eta_t'), skel0)
-     | _ => (case rews (sort_rrules (Net.match_term rules eta_t)) of
-               NONE => proc_rews (Net.match_term procs eta_t)
-             | some => some)
+  in
+    (case eta_t of
+      Abs _ $ _ => SOME (Thm.transitive eta_thm (Thm.beta_conversion false eta_t'), skel0)
+    | _ =>
+      (case rews (sort_rrules (Net.match_term rules eta_t)) of
+        NONE => proc_rews (Net.match_term procs eta_t)
+      | some => some))
   end;
 
 
@@ -991,13 +1002,15 @@
       val thm' = Thm.instantiate insts (Thm.rename_boundvars (term_of rlhs) (term_of t) rthm);
       val _ = trace_thm (fn () => "Applying congruence rule:") ss thm';
       fun err (msg, thm) = (trace_thm (fn () => msg) ss thm; NONE)
-  in case prover thm' of
-       NONE => err ("Congruence proof failed.  Could not prove", thm')
-     | SOME thm2 => (case check_conv true ss (Drule.beta_eta_conversion t) thm2 of
+  in
+    (case prover thm' of
+      NONE => err ("Congruence proof failed.  Could not prove", thm')
+    | SOME thm2 =>
+        (case check_conv true ss (Drule.beta_eta_conversion t) thm2 of
           NONE => err ("Congruence proof failed.  Should not have proved", thm2)
         | SOME thm2' =>
             if op aconv (pairself term_of (Thm.dest_equals (cprop_of thm2')))
-            then NONE else SOME thm2')
+            then NONE else SOME thm2'))
   end;
 
 val (cA, (cB, cC)) =
@@ -1141,19 +1154,21 @@
             val concl' =
               Drule.mk_implies (prem, the_default concl (Option.map Thm.rhs_of eq));
             val dprem = Option.map (disch false prem)
-          in case rewritec (prover, thy, maxidx) ss' concl' of
+          in
+            (case rewritec (prover, thy, maxidx) ss' concl' of
               NONE => rebuild prems concl' rrss asms ss (dprem eq)
             | SOME (eq', _) => transitive2 (fold (disch false)
                   prems (the (transitive3 (dprem eq) eq')))
-                (mut_impc0 (rev prems) (Thm.rhs_of eq') (rev rrss) (rev asms) ss)
+                (mut_impc0 (rev prems) (Thm.rhs_of eq') (rev rrss) (rev asms) ss))
           end
 
     and mut_impc0 prems concl rrss asms ss =
       let
         val prems' = strip_imp_prems concl;
         val (rrss', asms') = split_list (map (rules_of_prem ss) prems')
-      in mut_impc (prems @ prems') (strip_imp_concl concl) (rrss @ rrss')
-        (asms @ asms') [] [] [] [] ss ~1 ~1
+      in
+        mut_impc (prems @ prems') (strip_imp_concl concl) (rrss @ rrss')
+          (asms @ asms') [] [] [] [] ss ~1 ~1
       end
 
     and mut_impc [] concl [] [] prems' rrss' asms' eqns ss changed k =
@@ -1188,24 +1203,28 @@
             end
 
      (*legacy code - only for backwards compatibility*)
-     and nonmut_impc ct ss =
-       let val (prem, conc) = Thm.dest_implies ct;
-           val thm1 = if simprem then botc skel0 ss prem else NONE;
-           val prem1 = the_default prem (Option.map Thm.rhs_of thm1);
-           val ss1 = if not useprem then ss else add_rrules
-             (apsnd single (apfst single (rules_of_prem ss prem1))) ss
-       in (case botc skel0 ss1 conc of
-           NONE => (case thm1 of
-               NONE => NONE
-             | SOME thm1' => SOME (Drule.imp_cong_rule thm1' (Thm.reflexive conc)))
-         | SOME thm2 =>
-           let val thm2' = disch false prem1 thm2
-           in (case thm1 of
-               NONE => SOME thm2'
-             | SOME thm1' =>
+    and nonmut_impc ct ss =
+      let
+        val (prem, conc) = Thm.dest_implies ct;
+        val thm1 = if simprem then botc skel0 ss prem else NONE;
+        val prem1 = the_default prem (Option.map Thm.rhs_of thm1);
+        val ss1 =
+          if not useprem then ss
+          else add_rrules (apsnd single (apfst single (rules_of_prem ss prem1))) ss
+      in
+        (case botc skel0 ss1 conc of
+          NONE =>
+            (case thm1 of
+              NONE => NONE
+            | SOME thm1' => SOME (Drule.imp_cong_rule thm1' (Thm.reflexive conc)))
+        | SOME thm2 =>
+            let val thm2' = disch false prem1 thm2 in
+              (case thm1 of
+                NONE => SOME thm2'
+              | SOME thm1' =>
                  SOME (Thm.transitive (Drule.imp_cong_rule thm1' (Thm.reflexive conc)) thm2'))
-           end)
-       end
+            end)
+      end
 
  in try_botc end;
 
@@ -1247,7 +1266,7 @@
     val depth = simp_depth ss;
     val _ =
       if depth mod 20 = 0 then
-        warning ("Simplification depth " ^ string_of_int depth)
+        if_visible ss warning ("Simplification depth " ^ string_of_int depth)
       else ();
     val _ = trace_cterm false (fn () => "SIMPLIFIER INVOKED ON THE FOLLOWING TERM:") ss ct;
     val _ = check_bounds ss ct;
--- a/src/Pure/search.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/search.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -18,9 +18,8 @@
   val THEN_MAYBE': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic
   val DEPTH_SOLVE: tactic -> tactic
   val DEPTH_SOLVE_1: tactic -> tactic
-  val iter_deepen_limit: int Unsynchronized.ref
-  val THEN_ITER_DEEPEN: tactic -> (thm -> bool) -> (int -> tactic) -> tactic
-  val ITER_DEEPEN: (thm -> bool) -> (int -> tactic) -> tactic
+  val THEN_ITER_DEEPEN: int -> tactic -> (thm -> bool) -> (int -> tactic) -> tactic
+  val ITER_DEEPEN: int -> (thm -> bool) -> (int -> tactic) -> tactic
   val trace_DEEPEN: bool Unsynchronized.ref
   val DEEPEN: int * int -> (int -> int -> tactic) -> int -> int -> tactic
   val trace_BEST_FIRST: bool Unsynchronized.ref
@@ -117,21 +116,18 @@
     in  prune_aux (take (qs, []))  end;
 
 
-(*No known example (on 1-5-2007) needs even thirty*)
-val iter_deepen_limit = Unsynchronized.ref 50;
-
 (*Depth-first iterative deepening search for a state that satisfies satp
   tactic tac0 sets up the initial goal queue, while tac1 searches it.
   The solution sequence is redundant: the cutoff heuristic makes it impossible
   to suppress solutions arising from earlier searches, as the accumulated cost
   (k) can be wrong.*)
-fun THEN_ITER_DEEPEN tac0 satp tac1 = traced_tac (fn st =>
+fun THEN_ITER_DEEPEN lim tac0 satp tac1 = traced_tac (fn st =>
  let val countr = Unsynchronized.ref 0
      and tf = tracify trace_DEPTH_FIRST (tac1 1)
      and qs0 = tac0 st
      (*bnd = depth bound; inc = estimate of increment required next*)
      fun depth (bnd,inc) [] =
-          if bnd > !iter_deepen_limit then
+          if bnd > lim then
              (tracing (string_of_int (!countr) ^
                        " inferences so far.  Giving up at " ^ string_of_int bnd);
               NONE)
@@ -170,19 +166,19 @@
                end
   in depth (0,5) [] end);
 
-val ITER_DEEPEN = THEN_ITER_DEEPEN all_tac;
+fun ITER_DEEPEN lim = THEN_ITER_DEEPEN lim all_tac;
 
 
 (*Simple iterative deepening tactical.  It merely "deepens" any search tactic
   using increment "inc" up to limit "lim". *)
 val trace_DEEPEN = Unsynchronized.ref false;
 
-fun DEEPEN (inc,lim) tacf m i =
+fun DEEPEN (inc, lim) tacf m i =
   let
     fun dpn m st =
       st |>
        (if has_fewer_prems i st then no_tac
-        else if m>lim then
+        else if m > lim then
           (if !trace_DEEPEN then tracing "Search depth limit exceeded: giving up" else ();
             no_tac)
         else
--- a/src/Pure/variable.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Pure/variable.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -506,7 +506,7 @@
     val tfrees = map #1 extras |> sort_distinct string_ord;
     val frees = map #2 extras |> sort_distinct string_ord;
   in
-    if null extras then ()
+    if null extras orelse not (Context_Position.is_visible ctxt2) then ()
     else warning ("Introduced fixed type variable(s): " ^ commas tfrees ^ " in " ^
       space_implode " or " (map quote frees))
   end;
--- a/src/Tools/Code/code_eval.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/Code/code_eval.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -9,8 +9,6 @@
   val target: string
   val eval: string option -> string * (unit -> 'a) option Unsynchronized.ref
     -> ((term -> term) -> 'a -> 'a) -> theory -> term -> string list -> 'a
-  val evaluation_code: theory -> string -> string list -> string list
-    -> string * ((string * string) list * (string * string) list)
   val setup: theory -> theory
 end;
 
@@ -21,16 +19,12 @@
 
 val target = "Eval";
 
-val eval_struct_name = "Code";
-
-fun evaluation_code thy struct_name_hint tycos consts =
+fun evaluation_code thy module_name tycos consts =
   let
     val (consts', (naming, program)) = Code_Thingol.consts_program thy false consts;
     val tycos' = map (the o Code_Thingol.lookup_tyco naming) tycos;
-    val struct_name = if struct_name_hint = "" then eval_struct_name
-      else struct_name_hint;
-    val (ml_code, target_names) = Code_ML.evaluation_code_of thy target
-      struct_name naming program (consts' @ tycos');
+    val (ml_code, target_names) = Code_Target.produce_code_for thy
+      target NONE module_name [] naming program (consts' @ tycos');
     val (consts'', tycos'') = chop (length consts') target_names;
     val consts_map = map2 (fn const => fn NONE =>
         error ("Constant " ^ (quote o Code.string_of_const thy) const
@@ -57,11 +51,11 @@
           |> Graph.new_node (value_name,
               Code_Thingol.Fun (Term.dummy_patternN, ((([], ty), [(([], t), (NONE, true))]), NONE)))
           |> fold (curry Graph.add_edge value_name) deps;
-        val (value_code, [SOME value_name']) = Code_ML.evaluation_code_of thy
-          (the_default target some_target) "" naming program' [value_name];
-        val sml_code = "let\n" ^ value_code ^ "\nin " ^ value_name'
-          ^ space_implode " " (map (enclose "(" ")") args) ^ " end";
-      in ML_Context.evaluate ctxt false reff sml_code end;
+        val (program_code, [SOME value_name']) = Code_Target.produce_code_for thy
+          (the_default target some_target) NONE "Code" [] naming program' [value_name];
+        val value_code = space_implode " "
+          (value_name' :: map (enclose "(" ")") args);
+      in ML_Context.evaluate ctxt false reff (program_code, value_code) end;
   in Code_Thingol.dynamic_eval_value thy postproc evaluator t end;
 
 
@@ -69,26 +63,23 @@
 
 local
 
-structure CodeAntiqData = Proof_Data
+structure Code_Antiq_Data = Proof_Data
 (
-  type T = (string list * string list) * (bool * (string
-    * (string * ((string * string) list * (string * string) list)) lazy));
-  fun init _ = (([], []), (true, ("", Lazy.value ("", ([], [])))));
+  type T = (string list * string list) * (bool
+    * (string * ((string * string) list * (string * string) list)) lazy);
+  fun init _ = (([], []), (true, (Lazy.value ("", ([], [])))));
 );
 
-val is_first_occ = fst o snd o CodeAntiqData.get;
+val is_first_occ = fst o snd o Code_Antiq_Data.get;
 
 fun register_code new_tycos new_consts ctxt =
   let
-    val ((tycos, consts), (_, (struct_name, _))) = CodeAntiqData.get ctxt;
+    val ((tycos, consts), _) = Code_Antiq_Data.get ctxt;
     val tycos' = fold (insert (op =)) new_tycos tycos;
     val consts' = fold (insert (op =)) new_consts consts;
-    val (struct_name', ctxt') = if struct_name = ""
-      then ML_Antiquote.variant eval_struct_name ctxt
-      else (struct_name, ctxt);
-    val acc_code = Lazy.lazy
-      (fn () => evaluation_code (ProofContext.theory_of ctxt) eval_struct_name tycos' consts');
-  in CodeAntiqData.put ((tycos', consts'), (false, (struct_name', acc_code))) ctxt' end;
+    val acc_code = Lazy.lazy (fn () =>
+      evaluation_code (ProofContext.theory_of ctxt) "Code" tycos' consts');
+  in Code_Antiq_Data.put ((tycos', consts'), (false, acc_code)) ctxt end;
 
 fun register_const const = register_code [] [const];
 
@@ -99,11 +90,11 @@
 
 fun print_code is_first print_it ctxt =
   let
-    val (_, (_, (struct_code_name, acc_code))) = CodeAntiqData.get ctxt;
+    val (_, (_, acc_code)) = Code_Antiq_Data.get ctxt;
     val (ml_code, (tycos_map, consts_map)) = Lazy.force acc_code;
     val ml_code = if is_first then ml_code
       else "";
-    val all_struct_name = "Isabelle." ^ struct_code_name;
+    val all_struct_name = "Isabelle";
   in (ml_code, print_it all_struct_name tycos_map consts_map) end;
 
 in
@@ -143,34 +134,30 @@
           Code_Printer.concat [Code_Printer.enum "," "(" ")" (map (pr' Code_Printer.BR) tys), tyco']
   in
     thy
-    |> Code_Target.add_syntax_tyco target tyco (SOME (k, pr))
+    |> Code_Target.add_tyco_syntax target tyco (SOME (k, pr))
   end;
 
 fun add_eval_constr (const, const') thy =
   let
     val k = Code.args_number thy const;
     fun pr pr' fxy ts = Code_Printer.brackify fxy
-      (const' :: the_list (Code_ML.print_tuple pr' Code_Printer.BR (map fst ts)));
+      (const' :: the_list (Code_Printer.tuplify pr' Code_Printer.BR (map fst ts)));
   in
     thy
-    |> Code_Target.add_syntax_const target const (SOME (Code_Printer.simple_const_syntax (k, pr)))
+    |> Code_Target.add_const_syntax target const (SOME (Code_Printer.simple_const_syntax (k, pr)))
   end;
 
-fun add_eval_const (const, const') = Code_Target.add_syntax_const target
+fun add_eval_const (const, const') = Code_Target.add_const_syntax target
   const (SOME (Code_Printer.simple_const_syntax (0, (K o K o K) const')));
 
 fun process (code_body, (tyco_map, (constr_map, const_map))) module_name NONE thy =
-      let
-        val pr = Code_Printer.str o Long_Name.append module_name;
-      in
-        thy
-        |> Code_Target.add_reserved target module_name
-        |> Context.theory_map
-          (ML_Context.exec (fn () => ML_Context.eval_text true Position.none code_body))
-        |> fold (add_eval_tyco o apsnd pr) tyco_map
-        |> fold (add_eval_constr o apsnd pr) constr_map
-        |> fold (add_eval_const o apsnd pr) const_map
-      end
+      thy
+      |> Code_Target.add_reserved target module_name
+      |> Context.theory_map
+        (ML_Context.exec (fn () => ML_Context.eval_text true Position.none code_body))
+      |> fold (add_eval_tyco o apsnd Code_Printer.str) tyco_map
+      |> fold (add_eval_constr o apsnd Code_Printer.str) constr_map
+      |> fold (add_eval_const o apsnd Code_Printer.str) const_map
   | process (code_body, _) _ (SOME file_name) thy =
       let
         val preamble =
--- a/src/Tools/Code/code_haskell.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/Code/code_haskell.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -24,11 +24,11 @@
 
 (** Haskell serializer **)
 
-fun print_haskell_stmt labelled_name syntax_class syntax_tyco syntax_const
+fun print_haskell_stmt labelled_name class_syntax tyco_syntax const_syntax
     reserved deresolve contr_classparam_typs deriving_show =
   let
     val deresolve_base = Long_Name.base_name o deresolve;
-    fun class_name class = case syntax_class class
+    fun class_name class = case class_syntax class
      of NONE => deresolve class
       | SOME class => class;
     fun print_typcontext tyvars vs = case maps (fn (v, sort) => map (pair v) sort) vs
@@ -43,7 +43,7 @@
           (map (str o lookup_var tyvars) vnames) @ str "." @@ Pretty.brk 1;
     fun print_tyco_expr tyvars fxy (tyco, tys) =
       brackify fxy (str tyco :: map (print_typ tyvars BR) tys)
-    and print_typ tyvars fxy (tycoexpr as tyco `%% tys) = (case syntax_tyco tyco
+    and print_typ tyvars fxy (tycoexpr as tyco `%% tys) = (case tyco_syntax tyco
          of NONE => print_tyco_expr tyvars fxy (deresolve tyco, tys)
           | SOME (i, print) => print (print_typ tyvars) fxy tys)
       | print_typ tyvars fxy (ITyVar v) = (str o lookup_var tyvars) v;
@@ -72,7 +72,7 @@
           in brackets (str "\\" :: ps @ str "->" @@ print_term tyvars some_thm vars' NOBR t') end
       | print_term tyvars some_thm vars fxy (ICase (cases as (_, t0))) =
           (case Code_Thingol.unfold_const_app t0
-           of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c)
+           of SOME (c_ts as ((c, _), _)) => if is_none (const_syntax c)
                 then print_case tyvars some_thm vars fxy cases
                 else print_app tyvars some_thm vars fxy c_ts
             | NONE => print_case tyvars some_thm vars fxy cases)
@@ -90,7 +90,7 @@
             (str o deresolve) c :: map2 print_term_anno ts_fingerprint (take (length ts) function_typs)
           else (str o deresolve) c :: map (print_term tyvars some_thm vars BR) ts
         end
-    and print_app tyvars = gen_print_app (print_app_expr tyvars) (print_term tyvars) syntax_const
+    and print_app tyvars = gen_print_app (print_app_expr tyvars) (print_term tyvars) const_syntax
     and print_bind tyvars some_thm fxy p = gen_print_bind (print_term tyvars) some_thm fxy p
     and print_case tyvars some_thm vars fxy (cases as ((_, [_]), _)) =
           let
@@ -133,7 +133,7 @@
                 val consts = fold Code_Thingol.add_constnames (t :: ts) [];
                 val vars = reserved
                   |> intro_base_names
-                      (is_none o syntax_const) deresolve consts
+                      (is_none o const_syntax) deresolve consts
                   |> intro_vars ((fold o Code_Thingol.fold_varnames)
                       (insert (op =)) ts []);
               in
@@ -218,7 +218,7 @@
       | print_stmt (_, Code_Thingol.Classinst ((class, (tyco, vs)), (_, (classparam_instances, _)))) =
           let
             val tyvars = intro_vars (map fst vs) reserved;
-            fun requires_args classparam = case syntax_const classparam
+            fun requires_args classparam = case const_syntax classparam
              of NONE => 0
               | SOME (Code_Printer.Plain_const_syntax _) => 0
               | SOME (Code_Printer.Complex_const_syntax (k,_ )) => k;
@@ -234,7 +234,7 @@
                       val (c, (_, tys)) = const;
                       val (vs, rhs) = (apfst o map) fst
                         (Code_Thingol.unfold_abs (Code_Thingol.eta_expand k (const, [])));
-                      val s = if (is_some o syntax_const) c
+                      val s = if (is_some o const_syntax) c
                         then NONE else (SOME o Long_Name.base_name o deresolve) c;
                       val vars = reserved
                         |> intro_vars (map_filter I (s :: vs));
@@ -261,9 +261,8 @@
           end;
   in print_stmt end;
 
-fun haskell_program_of_program labelled_name module_name module_prefix reserved raw_module_alias program =
+fun haskell_program_of_program labelled_name module_prefix reserved module_alias program =
   let
-    val module_alias = if is_some module_name then K module_name else raw_module_alias;
     val reserved = Name.make_context reserved;
     val mk_name_module = mk_name_module reserved module_prefix module_alias program;
     fun add_stmt (name, (stmt, deps)) =
@@ -314,15 +313,14 @@
       handle Option => error ("Unknown statement name: " ^ labelled_name name);
   in (deresolver, hs_program) end;
 
-fun serialize_haskell module_prefix raw_module_name string_classes labelled_name
-    raw_reserved includes raw_module_alias
-    syntax_class syntax_tyco syntax_const (code_of_pretty, code_writeln) program stmt_names destination =
+fun serialize_haskell module_prefix string_classes { labelled_name,
+    reserved_syms, includes, module_alias,
+    class_syntax, tyco_syntax, const_syntax, program,
+    names, presentation_names } =
   let
-    val presentation_stmt_names = Code_Target.stmt_names_of_destination destination;
-    val module_name = if null presentation_stmt_names then raw_module_name else SOME "Code";
-    val reserved = fold (insert (op =) o fst) includes raw_reserved;
+    val reserved = fold (insert (op =) o fst) includes reserved_syms;
     val (deresolver, hs_program) = haskell_program_of_program labelled_name
-      module_name module_prefix reserved raw_module_alias program;
+      module_prefix reserved module_alias program;
     val contr_classparam_typs = Code_Thingol.contr_classparam_typs program;
     fun deriving_show tyco =
       let
@@ -339,22 +337,20 @@
       in deriv [] tyco end;
     val reserved = make_vars reserved;
     fun print_stmt qualified = print_haskell_stmt labelled_name
-      syntax_class syntax_tyco syntax_const reserved
+      class_syntax tyco_syntax const_syntax reserved
       (if qualified then deresolver else Long_Name.base_name o deresolver)
       contr_classparam_typs
       (if string_classes then deriving_show else K false);
     fun print_module name content =
-      (name, Pretty.chunks [
+      (name, Pretty.chunks2 [
         str ("module " ^ name ^ " where {"),
-        str "",
         content,
-        str "",
         str "}"
       ]);
     fun serialize_module1 (module_name', (deps, (stmts, _))) =
       let
         val stmt_names = map fst stmts;
-        val qualified = is_none module_name;
+        val qualified = null presentation_names;
         val imports = subtract (op =) stmt_names deps
           |> distinct (op =)
           |> map_filter (try deresolver)
@@ -372,37 +368,41 @@
           );
       in print_module module_name' content end;
     fun serialize_module2 (_, (_, (stmts, _))) = Pretty.chunks2 (map_filter
-        (fn (name, (_, SOME stmt)) => if null presentation_stmt_names
-              orelse member (op =) presentation_stmt_names name
+        (fn (name, (_, SOME stmt)) => if null presentation_names
+              orelse member (op =) presentation_names name
               then SOME (print_stmt false (name, stmt))
               else NONE
           | (_, (_, NONE)) => NONE) stmts);
     val serialize_module =
-      if null presentation_stmt_names then serialize_module1 else pair "" o serialize_module2;
-    fun check_destination destination =
-      (File.check destination; destination);
-    fun write_module destination (modlname, content) =
-      let
-        val filename = case modlname
-         of "" => Path.explode "Main.hs"
-          | _ => (Path.ext "hs" o Path.explode o implode o separate "/"
-                o Long_Name.explode) modlname;
-        val pathname = Path.append destination filename;
-        val _ = File.mkdir_leaf (Path.dir pathname);
-      in File.write pathname
-        ("{-# OPTIONS_GHC -fglasgow-exts #-}\n\n"
-          ^ code_of_pretty content)
-      end
+      if null presentation_names then serialize_module1 else pair "" o serialize_module2;
+    fun write_module width (SOME destination) (modlname, content) =
+          let
+            val _ = File.check destination;
+            val filename = case modlname
+             of "" => Path.explode "Main.hs"
+              | _ => (Path.ext "hs" o Path.explode o implode o separate "/"
+                    o Long_Name.explode) modlname;
+            val pathname = Path.append destination filename;
+            val _ = File.mkdir_leaf (Path.dir pathname);
+          in File.write pathname
+            ("{-# OPTIONS_GHC -fglasgow-exts #-}\n\n"
+              ^ string_of_pretty width content)
+          end
+      | write_module width NONE (_, content) = writeln_pretty width content;
   in
-    Code_Target.mk_serialization target
-      (fn NONE => K () o map (code_writeln o (fn p => Pretty.block [p, Pretty.fbrk]) o snd)
-        | SOME file => K () o map (write_module (check_destination file)))
-      (rpair [] o cat_lines o map (code_of_pretty o snd))
+    Code_Target.serialization
+      (fn width => fn destination => K () o map (write_module width destination))
+      (fn width => rpair [] o cat_lines o map (string_of_pretty width o snd))
       (map (uncurry print_module) includes
         @ map serialize_module (Symtab.dest hs_program))
-      destination
   end;
 
+val serializer : Code_Target.serializer =
+  Code_Target.parse_args (Scan.option (Args.$$$ "root" -- Args.colon |-- Args.name)
+    -- Scan.optional (Args.$$$ "string_classes" >> K true) false
+    >> (fn (module_prefix, string_classes) =>
+      serialize_haskell module_prefix string_classes));
+
 val literals = let
   fun char_haskell c =
     let
@@ -464,19 +464,13 @@
     val c_bind = Code.read_const thy raw_c_bind;
   in if target = target' then
     thy
-    |> Code_Target.add_syntax_const target c_bind
+    |> Code_Target.add_const_syntax target c_bind
         (SOME (Code_Printer.complex_const_syntax (pretty_haskell_monad c_bind)))
   else error "Only Haskell target allows for monad syntax" end;
 
 
 (** Isar setup **)
 
-fun isar_serializer module_name =
-  Code_Target.parse_args (Scan.option (Args.$$$ "root" -- Args.colon |-- Args.name)
-    -- Scan.optional (Args.$$$ "string_classes" >> K true) false
-    >> (fn (module_prefix, string_classes) =>
-      serialize_haskell module_prefix module_name string_classes));
-
 val _ =
   Outer_Syntax.command "code_monad" "define code syntax for monads" Keyword.thy_decl (
     Parse.term_group -- Parse.name >> (fn (raw_bind, target) =>
@@ -485,11 +479,11 @@
 
 val setup =
   Code_Target.add_target
-    (target, { serializer = isar_serializer, literals = literals,
+    (target, { serializer = serializer, literals = literals,
       check = { env_var = "EXEC_GHC", make_destination = I,
-        make_command = fn ghc => fn p => fn module_name =>
+        make_command = fn ghc => fn module_name =>
           ghc ^ " -fglasgow-exts -odir build -hidir build -stubdir build -e \"\" " ^ module_name ^ ".hs" } })
-  #> Code_Target.add_syntax_tyco target "fun" (SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>
+  #> Code_Target.add_tyco_syntax target "fun" (SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>
       brackify_infix (1, R) fxy (
         print_typ (INFX (1, X)) ty1,
         str "->",
--- a/src/Tools/Code/code_ml.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/Code/code_ml.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -8,10 +8,6 @@
 sig
   val target_SML: string
   val target_OCaml: string
-  val evaluation_code_of: theory -> string -> string
-    -> Code_Thingol.naming -> Code_Thingol.program -> string list -> string * string option list
-  val print_tuple: (Code_Printer.fixity -> 'a -> Pretty.T)
-    -> Code_Printer.fixity -> 'a list -> Pretty.T option
   val setup: theory -> theory
 end;
 
@@ -56,21 +52,21 @@
   | print_product print [x] = SOME (print x)
   | print_product print xs = (SOME o enum " *" "" "") (map print xs);
 
-fun print_tuple _ _ [] = NONE
-  | print_tuple print fxy [x] = SOME (print fxy x)
-  | print_tuple print _ xs = SOME (enum "," "(" ")" (map (print NOBR) xs));
+fun tuplify _ _ [] = NONE
+  | tuplify print fxy [x] = SOME (print fxy x)
+  | tuplify print _ xs = SOME (enum "," "(" ")" (map (print NOBR) xs));
 
 
 (** SML serializer **)
 
-fun print_sml_stmt labelled_name syntax_tyco syntax_const reserved is_cons deresolve =
+fun print_sml_stmt labelled_name tyco_syntax const_syntax reserved is_cons deresolve =
   let
     fun print_tyco_expr fxy (tyco, []) = (str o deresolve) tyco
       | print_tyco_expr fxy (tyco, [ty]) =
           concat [print_typ BR ty, (str o deresolve) tyco]
       | print_tyco_expr fxy (tyco, tys) =
           concat [enum "," "(" ")" (map (print_typ BR) tys), (str o deresolve) tyco]
-    and print_typ fxy (tyco `%% tys) = (case syntax_tyco tyco
+    and print_typ fxy (tyco `%% tys) = (case tyco_syntax tyco
          of NONE => print_tyco_expr fxy (tyco, tys)
           | SOME (i, print) => print print_typ fxy tys)
       | print_typ fxy (ITyVar v) = str ("'" ^ v);
@@ -94,7 +90,7 @@
             | classrels => brackets
                 [enum " o" "(" ")" (map (str o deresolve) classrels), v_p]
           end
-    and print_dicts is_pseudo_fun = print_tuple (print_dict is_pseudo_fun);
+    and print_dicts is_pseudo_fun = tuplify (print_dict is_pseudo_fun);
     val print_dict_args = map_filter (fn (v, sort) => print_dicts (K false) BR
       (map_index (fn (i, _) => DictVar ([], (v, (i, length sort)))) sort));
     fun print_term is_pseudo_fun some_thm vars fxy (IConst c) =
@@ -118,7 +114,7 @@
           in brackets (ps @ [print_term is_pseudo_fun some_thm vars' NOBR t']) end
       | print_term is_pseudo_fun some_thm vars fxy (ICase (cases as (_, t0))) =
           (case Code_Thingol.unfold_const_app t0
-           of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c)
+           of SOME (c_ts as ((c, _), _)) => if is_none (const_syntax c)
                 then print_case is_pseudo_fun some_thm vars fxy cases
                 else print_app is_pseudo_fun some_thm vars fxy c_ts
             | NONE => print_case is_pseudo_fun some_thm vars fxy cases)
@@ -127,7 +123,7 @@
         let val k = length function_typs in
           if k < 2 orelse length ts = k
           then (str o deresolve) c
-            :: the_list (print_tuple (print_term is_pseudo_fun some_thm vars) BR ts)
+            :: the_list (tuplify (print_term is_pseudo_fun some_thm vars) BR ts)
           else [print_term is_pseudo_fun some_thm vars BR (Code_Thingol.eta_expand k app)]
         end
       else if is_pseudo_fun c
@@ -135,7 +131,7 @@
       else (str o deresolve) c :: map_filter (print_dicts is_pseudo_fun BR) iss
         @ map (print_term is_pseudo_fun some_thm vars BR) ts
     and print_app is_pseudo_fun some_thm vars = gen_print_app (print_app_expr is_pseudo_fun)
-      (print_term is_pseudo_fun) syntax_const some_thm vars
+      (print_term is_pseudo_fun) const_syntax some_thm vars
     and print_bind is_pseudo_fun = gen_print_bind (print_term is_pseudo_fun)
     and print_case is_pseudo_fun some_thm vars fxy (cases as ((_, [_]), _)) =
           let
@@ -194,7 +190,7 @@
                 val consts = fold Code_Thingol.add_constnames (t :: ts) [];
                 val vars = reserved
                   |> intro_base_names
-                       (is_none o syntax_const) deresolve consts
+                       (is_none o const_syntax) deresolve consts
                   |> intro_vars ((fold o Code_Thingol.fold_varnames)
                        (insert (op =)) ts []);
                 val prolog = if needs_typ then
@@ -343,9 +339,8 @@
           end;
   in print_stmt end;
 
-fun print_sml_module name some_decls body = if name = ""
-  then Pretty.chunks2 body
-  else Pretty.chunks2 (
+fun print_sml_module name some_decls body =
+  Pretty.chunks2 (
     Pretty.chunks (
       str ("structure " ^ name ^ (if is_some some_decls then " : sig" else " ="))
       :: (the_list o Option.map (indent 2 o Pretty.chunks)) some_decls
@@ -369,14 +364,14 @@
 
 (** OCaml serializer **)
 
-fun print_ocaml_stmt labelled_name syntax_tyco syntax_const reserved is_cons deresolve =
+fun print_ocaml_stmt labelled_name tyco_syntax const_syntax reserved is_cons deresolve =
   let
     fun print_tyco_expr fxy (tyco, []) = (str o deresolve) tyco
       | print_tyco_expr fxy (tyco, [ty]) =
           concat [print_typ BR ty, (str o deresolve) tyco]
       | print_tyco_expr fxy (tyco, tys) =
           concat [enum "," "(" ")" (map (print_typ BR) tys), (str o deresolve) tyco]
-    and print_typ fxy (tyco `%% tys) = (case syntax_tyco tyco
+    and print_typ fxy (tyco `%% tys) = (case tyco_syntax tyco
          of NONE => print_tyco_expr fxy (tyco, tys)
           | SOME (i, print) => print print_typ fxy tys)
       | print_typ fxy (ITyVar v) = str ("'" ^ v);
@@ -395,7 +390,7 @@
             else "_" ^ first_upper v ^ string_of_int (i+1))
           |> fold_rev (fn classrel => fn p =>
                Pretty.block [p, str ".", (str o deresolve) classrel]) classrels
-    and print_dicts is_pseudo_fun = print_tuple (print_dict is_pseudo_fun);
+    and print_dicts is_pseudo_fun = tuplify (print_dict is_pseudo_fun);
     val print_dict_args = map_filter (fn (v, sort) => print_dicts (K false) BR
       (map_index (fn (i, _) => DictVar ([], (v, (i, length sort)))) sort));
     fun print_term is_pseudo_fun some_thm vars fxy (IConst c) =
@@ -416,7 +411,7 @@
           in brackets (str "fun" :: ps @ str "->" @@ print_term is_pseudo_fun some_thm vars' NOBR t') end
       | print_term is_pseudo_fun some_thm vars fxy (ICase (cases as (_, t0))) =
           (case Code_Thingol.unfold_const_app t0
-           of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c)
+           of SOME (c_ts as ((c, _), _)) => if is_none (const_syntax c)
                 then print_case is_pseudo_fun some_thm vars fxy cases
                 else print_app is_pseudo_fun some_thm vars fxy c_ts
             | NONE => print_case is_pseudo_fun some_thm vars fxy cases)
@@ -425,7 +420,7 @@
         let val k = length tys in
           if length ts = k
           then (str o deresolve) c
-            :: the_list (print_tuple (print_term is_pseudo_fun some_thm vars) BR ts)
+            :: the_list (tuplify (print_term is_pseudo_fun some_thm vars) BR ts)
           else [print_term is_pseudo_fun some_thm vars BR (Code_Thingol.eta_expand k app)]
         end
       else if is_pseudo_fun c
@@ -433,7 +428,7 @@
       else (str o deresolve) c :: map_filter (print_dicts is_pseudo_fun BR) iss
         @ map (print_term is_pseudo_fun some_thm vars BR) ts
     and print_app is_pseudo_fun some_thm vars = gen_print_app (print_app_expr is_pseudo_fun)
-      (print_term is_pseudo_fun) syntax_const some_thm vars
+      (print_term is_pseudo_fun) const_syntax some_thm vars
     and print_bind is_pseudo_fun = gen_print_bind (print_term is_pseudo_fun)
     and print_case is_pseudo_fun some_thm vars fxy (cases as ((_, [_]), _)) =
           let
@@ -489,7 +484,7 @@
                   |> intro_vars ((fold o Code_Thingol.fold_varnames)
                       (insert (op =)) ts []);
               in concat [
-                (Pretty.block o Pretty.commas)
+                (Pretty.block o commas)
                   (map (print_term is_pseudo_fun some_thm vars NOBR) ts),
                 str "->",
                 print_term is_pseudo_fun some_thm vars NOBR t
@@ -499,7 +494,7 @@
                     val consts = fold Code_Thingol.add_constnames (t :: ts) [];
                     val vars = reserved
                       |> intro_base_names
-                          (is_none o syntax_const) deresolve consts
+                          (is_none o const_syntax) deresolve consts
                       |> intro_vars ((fold o Code_Thingol.fold_varnames)
                           (insert (op =)) ts []);
                   in
@@ -525,7 +520,7 @@
                     val consts = fold Code_Thingol.add_constnames (map (snd o fst) eqs) [];
                     val vars = reserved
                       |> intro_base_names
-                          (is_none o syntax_const) deresolve consts;
+                          (is_none o const_syntax) deresolve consts;
                     val dummy_parms = (map str o aux_params vars o map (fst o fst)) eqs;
                   in
                     Pretty.block (
@@ -535,7 +530,7 @@
                       :: Pretty.brk 1
                       :: str "match"
                       :: Pretty.brk 1
-                      :: (Pretty.block o Pretty.commas) dummy_parms
+                      :: (Pretty.block o commas) dummy_parms
                       :: Pretty.brk 1
                       :: str "with"
                       :: Pretty.brk 1
@@ -669,9 +664,8 @@
           end;
   in print_stmt end;
 
-fun print_ocaml_module name some_decls body = if name = ""
-  then Pretty.chunks2 body
-  else Pretty.chunks2 (
+fun print_ocaml_module name some_decls body =
+  Pretty.chunks2 (
     Pretty.chunks (
       str ("module " ^ name ^ (if is_some some_decls then " : sig" else " ="))
       :: (the_list o Option.map (indent 2 o Pretty.chunks)) some_decls
@@ -722,9 +716,8 @@
 
 in
 
-fun ml_node_of_program labelled_name module_name reserved raw_module_alias program =
+fun ml_node_of_program labelled_name reserved module_alias program =
   let
-    val module_alias = if is_some module_name then K module_name else raw_module_alias;
     val reserved = Name.make_context reserved;
     val empty_module = ((reserved, reserved), Graph.empty);
     fun map_node [] f = f
@@ -813,7 +806,7 @@
         ) stmts
       #>> (split_list #> apsnd (map_filter I
         #> (fn [] => error ("Datatype block without data statement: "
-                  ^ (commas o map (labelled_name o fst)) stmts)
+                  ^ (Library.commas o map (labelled_name o fst)) stmts)
              | stmts => ML_Datas stmts)));
     fun add_class stmts =
       fold_map
@@ -828,7 +821,7 @@
         ) stmts
       #>> (split_list #> apsnd (map_filter I
         #> (fn [] => error ("Class block without class statement: "
-                  ^ (commas o map (labelled_name o fst)) stmts)
+                  ^ (Library.commas o map (labelled_name o fst)) stmts)
              | [stmt] => ML_Class stmt)));
     fun add_stmts ([stmt as (name, Code_Thingol.Fun _)]) =
           add_fun stmt
@@ -849,7 +842,7 @@
       | add_stmts ((stmts as (_, Code_Thingol.Classinst _)::_)) =
           add_funs stmts
       | add_stmts stmts = error ("Illegal mutual dependencies: " ^
-          (commas o map (labelled_name o fst)) stmts);
+          (Library.commas o map (labelled_name o fst)) stmts);
     fun add_stmts' stmts nsp_nodes =
       let
         val names as (name :: names') = map fst stmts;
@@ -858,9 +851,9 @@
         val module_name = (the_single o distinct (op =) o map mk_name_module) module_names
           handle Empty =>
             error ("Different namespace prefixes for mutual dependencies:\n"
-              ^ commas (map labelled_name names)
+              ^ Library.commas (map labelled_name names)
               ^ "\n"
-              ^ commas module_names);
+              ^ Library.commas module_names);
         val module_name_path = Long_Name.explode module_name;
         fun add_dep name name' =
           let
@@ -907,22 +900,21 @@
         error ("Unknown statement name: " ^ labelled_name name);
   in (deresolver, nodes) end;
 
-fun serialize_ml target print_module print_stmt raw_module_name with_signatures labelled_name
-  reserved includes raw_module_alias _ syntax_tyco syntax_const (code_of_pretty, code_writeln) program stmt_names destination =
+fun serialize_ml target print_module print_stmt with_signatures { labelled_name,
+  reserved_syms, includes, module_alias, class_syntax, tyco_syntax,
+  const_syntax, program, names, presentation_names } =
   let
     val is_cons = Code_Thingol.is_cons program;
-    val presentation_stmt_names = Code_Target.stmt_names_of_destination destination;
-    val is_presentation = not (null presentation_stmt_names);
-    val module_name = if is_presentation then SOME "Code" else raw_module_name;
-    val (deresolver, nodes) = ml_node_of_program labelled_name module_name
-      reserved raw_module_alias program;
-    val reserved = make_vars reserved;
+    val is_presentation = not (null presentation_names);
+    val (deresolver, nodes) = ml_node_of_program labelled_name
+      reserved_syms module_alias program;
+    val reserved = make_vars reserved_syms;
     fun print_node prefix (Dummy _) =
           NONE
       | print_node prefix (Stmt (_, stmt)) = if is_presentation andalso
-          (null o filter (member (op =) presentation_stmt_names) o stmt_names_of) stmt
+          (null o filter (member (op =) presentation_names) o stmt_names_of) stmt
           then NONE
-          else SOME (print_stmt labelled_name syntax_tyco syntax_const reserved is_cons (deresolver prefix) stmt)
+          else SOME (print_stmt labelled_name tyco_syntax const_syntax reserved is_cons (deresolver prefix) stmt)
       | print_node prefix (Module (module_name, (_, nodes))) =
           let
             val (decls, body) = print_nodes (prefix @ [module_name]) nodes;
@@ -933,53 +925,45 @@
         o rev o flat o Graph.strong_conn) nodes
       |> split_list
       |> (fn (decls, body) => (flat decls, body))
-    val stmt_names' = (map o try)
-      (deresolver (if is_some module_name then the_list module_name else [])) stmt_names;
+    val names' = map (try (deresolver [])) names;
     val p = Pretty.chunks2 (map snd includes @ snd (print_nodes [] nodes));
+    fun write width NONE = writeln_pretty width
+      | write width (SOME p) = File.write p o string_of_pretty width;
   in
-    Code_Target.mk_serialization target
-      (fn NONE => code_writeln | SOME file => File.write file o code_of_pretty)
-      (rpair stmt_names' o code_of_pretty) p destination
+    Code_Target.serialization write (fn width => (rpair names' o string_of_pretty width)) p
   end;
 
 end; (*local*)
 
-
-(** for instrumentalization **)
+val serializer_sml : Code_Target.serializer =
+  Code_Target.parse_args (Scan.optional (Args.$$$ "no_signatures" >> K false) true
+  >> (fn with_signatures => serialize_ml target_SML
+      print_sml_module print_sml_stmt with_signatures));
 
-fun evaluation_code_of thy target struct_name =
-  Code_Target.serialize_custom thy (target, fn _ => fn [] =>
-    serialize_ml target print_sml_module print_sml_stmt (SOME struct_name) true);
+val serializer_ocaml : Code_Target.serializer =
+  Code_Target.parse_args (Scan.optional (Args.$$$ "no_signatures" >> K false) true
+  >> (fn with_signatures => serialize_ml target_OCaml
+      print_ocaml_module print_ocaml_stmt with_signatures));
 
 
 (** Isar setup **)
 
-fun isar_serializer_sml module_name =
-  Code_Target.parse_args (Scan.optional (Args.$$$ "no_signatures" >> K false) true
-  >> (fn with_signatures => serialize_ml target_SML
-      print_sml_module print_sml_stmt module_name with_signatures));
-
-fun isar_serializer_ocaml module_name =
-  Code_Target.parse_args (Scan.optional (Args.$$$ "no_signatures" >> K false) true
-  >> (fn with_signatures => serialize_ml target_OCaml
-      print_ocaml_module print_ocaml_stmt module_name with_signatures));
-
 val setup =
   Code_Target.add_target
-    (target_SML, { serializer = isar_serializer_sml, literals = literals_sml,
+    (target_SML, { serializer = serializer_sml, literals = literals_sml,
       check = { env_var = "ISABELLE_PROCESS", make_destination = fn p => Path.append p (Path.explode "ROOT.ML"),
-        make_command = fn isabelle => fn p => fn _ => isabelle ^ " -r -q -u Pure" } })
+        make_command = fn isabelle => fn _ => isabelle ^ " -r -q -u Pure" } })
   #> Code_Target.add_target
-    (target_OCaml, { serializer = isar_serializer_ocaml, literals = literals_ocaml,
+    (target_OCaml, { serializer = serializer_ocaml, literals = literals_ocaml,
       check = { env_var = "EXEC_OCAML", make_destination = fn p => Path.append p (Path.explode "ROOT.ocaml"),
-        make_command = fn ocaml => fn p => fn _ => ocaml ^ " -w pu nums.cma " ^ File.shell_path p } })
-  #> Code_Target.add_syntax_tyco target_SML "fun" (SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>
+        make_command = fn ocaml => fn _ => ocaml ^ " -w pu nums.cma ROOT.ocaml" } })
+  #> Code_Target.add_tyco_syntax target_SML "fun" (SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>
       brackify_infix (1, R) fxy (
         print_typ (INFX (1, X)) ty1,
         str "->",
         print_typ (INFX (1, R)) ty2
       )))
-  #> Code_Target.add_syntax_tyco target_OCaml "fun" (SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>
+  #> Code_Target.add_tyco_syntax target_OCaml "fun" (SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>
       brackify_infix (1, R) fxy (
         print_typ (INFX (1, X)) ty1,
         str "->",
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Code/code_namespace.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -0,0 +1,129 @@
+(*  Title:      Tools/Code/code_namespace.ML
+    Author:     Florian Haftmann, TU Muenchen
+
+Mastering target language namespaces.
+*)
+
+signature CODE_NAMESPACE =
+sig
+  datatype 'a node =
+      Dummy
+    | Stmt of Code_Thingol.stmt
+    | Module of ('a * (string * 'a node) Graph.T);
+  val hierarchical_program: (string -> string) -> { module_alias: string -> string option,
+    reserved: Name.context, empty_nsp: 'b, namify_module: string -> 'b -> string * 'b,
+    namify_stmt: Code_Thingol.stmt -> string -> 'b -> string * 'b,
+    cyclic_modules: bool, empty_data: 'a, memorize_data: string -> 'a -> 'a }
+      -> Code_Thingol.program
+      -> { deresolver: string list -> string -> string,
+           hierarchical_program: (string * 'a node) Graph.T }
+end;
+
+structure Code_Namespace : CODE_NAMESPACE =
+struct
+
+(* hierarchical program structure *)
+
+datatype 'a node =
+    Dummy
+  | Stmt of Code_Thingol.stmt
+  | Module of ('a * (string * 'a node) Graph.T);
+
+fun hierarchical_program labelled_name { module_alias, reserved, empty_nsp,
+      namify_module, namify_stmt, cyclic_modules, empty_data, memorize_data } program =
+  let
+
+    (* building module name hierarchy *)
+    fun alias_fragments name = case module_alias name
+     of SOME name' => Long_Name.explode name'
+      | NONE => map (fn name => fst (yield_singleton Name.variants name reserved))
+          (Long_Name.explode name);
+    val module_names = Graph.fold (insert (op =) o fst o Code_Printer.dest_name o fst) program [];
+    val fragments_tab = fold (fn name => Symtab.update
+      (name, alias_fragments name)) module_names Symtab.empty;
+    val dest_name = Code_Printer.dest_name #>> (the o Symtab.lookup fragments_tab);
+
+    (* building empty module hierarchy *)
+    val empty_module = (empty_data, Graph.empty);
+    fun map_module f (Module content) = Module (f content);
+    fun change_module [] = I
+      | change_module (name_fragment :: name_fragments) =
+          apsnd o Graph.map_node name_fragment o apsnd o map_module
+            o change_module name_fragments;
+    fun ensure_module name_fragment (data, nodes) =
+      if can (Graph.get_node nodes) name_fragment then (data, nodes)
+      else (data,
+        nodes |> Graph.new_node (name_fragment, (name_fragment, Module empty_module)));
+    fun allocate_module [] = I
+      | allocate_module (name_fragment :: name_fragments) =
+          ensure_module name_fragment
+          #> (apsnd o Graph.map_node name_fragment o apsnd o map_module o allocate_module) name_fragments;
+    val empty_program = Symtab.fold (fn (_, fragments) => allocate_module fragments)
+      fragments_tab empty_module;
+
+    (* distribute statements over hierarchy *)
+    fun add_stmt name stmt =
+      let
+        val (name_fragments, base) = dest_name name;
+      in
+        change_module name_fragments (fn (data, nodes) =>
+          (memorize_data name data, Graph.new_node (name, (base, Stmt stmt)) nodes))
+      end;
+    fun add_dependency name name' =
+      let
+        val (name_fragments, base) = dest_name name;
+        val (name_fragments', base') = dest_name name';
+        val (name_fragments_common, (diff, diff')) =
+          chop_prefix (op =) (name_fragments, name_fragments');
+        val (is_module, dep) = if null diff then (false, (name, name'))
+          else (true, (hd diff, hd diff'))
+        val add_edge = if is_module andalso not cyclic_modules
+          then (fn node => Graph.add_edge_acyclic dep node
+            handle Graph.CYCLES _ => error ("Dependency "
+              ^ quote name ^ " -> " ^ quote name'
+              ^ " would result in module dependency cycle"))
+          else Graph.add_edge dep
+      in (change_module name_fragments_common o apsnd) add_edge end;
+    val proto_program = empty_program
+      |> Graph.fold (fn (name, (stmt, _)) => add_stmt name stmt) program
+      |> Graph.fold (fn (name, (_, (_, names))) => fold (add_dependency name) names) program;
+
+    (* name declarations *)
+    fun make_declarations nsps (data, nodes) =
+      let
+        val (module_fragments, stmt_names) = List.partition
+          (fn name_fragment => case Graph.get_node nodes name_fragment
+            of (_, Module _) => true | _ => false) (Graph.keys nodes);
+        fun modify_stmt (Stmt (Code_Thingol.Datatypecons _)) = Dummy
+          | modify_stmt (Stmt (Code_Thingol.Classrel _)) = Dummy
+          | modify_stmt (Stmt (Code_Thingol.Classparam _)) = Dummy
+          | modify_stmt stmt = stmt;
+        fun declare namify modify name (nsps, nodes) =
+          let
+            val (base, node) = Graph.get_node nodes name;
+            val (base', nsps') = namify node base nsps;
+            val nodes' = Graph.map_node name (K (base', modify node)) nodes;
+          in (nsps', nodes') end;
+        val (nsps', nodes') = (nsps, nodes)
+          |> fold (declare (K namify_module) I) module_fragments
+          |> fold (declare (namify_stmt o (fn Stmt stmt => stmt)) modify_stmt) stmt_names;
+        val nodes'' = nodes'
+          |> fold (fn name_fragment => (Graph.map_node name_fragment
+              o apsnd o map_module) (make_declarations nsps')) module_fragments;
+      in (data, nodes'') end;
+    val (_, hierarchical_program) = make_declarations empty_nsp proto_program;
+
+    (* deresolving *)
+    fun deresolver prefix_fragments name =
+      let
+        val (name_fragments, _) = dest_name name;
+        val (_, (_, remainder)) = chop_prefix (op =) (prefix_fragments, name_fragments);
+        val nodes = fold (fn name_fragment => fn nodes => case Graph.get_node nodes name_fragment
+         of (_, Module (_, nodes)) => nodes) name_fragments hierarchical_program;
+        val (base', _) = Graph.get_node nodes name;
+      in Long_Name.implode (remainder @ [base']) end
+        handle Graph.UNDEF _ => error ("Unknown statement name: " ^ labelled_name name);
+
+  in { deresolver = deresolver, hierarchical_program = hierarchical_program } end;
+
+end;
\ No newline at end of file
--- a/src/Tools/Code/code_preproc.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/Code/code_preproc.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -92,7 +92,7 @@
     val thm_sym = Thm.symmetric thm;
   in
     thy |> map_pre_post (fn (pre, post) =>
-      (pre |> MetaSimplifier.add_simp thm, post |> MetaSimplifier.del_simp thm_sym))
+      (pre |> MetaSimplifier.add_simp thm, post |> MetaSimplifier.add_simp thm_sym))
   end;
 
 fun add_functrans (name, f) = (map_data o apsnd)
--- a/src/Tools/Code/code_printer.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/Code/code_printer.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -19,6 +19,7 @@
   val concat: Pretty.T list -> Pretty.T
   val brackets: Pretty.T list -> Pretty.T
   val enclose: string -> string -> Pretty.T list -> Pretty.T
+  val commas: Pretty.T list -> Pretty.T list
   val enum: string -> string -> string -> Pretty.T list -> Pretty.T
   val enum_default: string -> string -> string -> string -> Pretty.T list -> Pretty.T
   val semicolon: Pretty.T list -> Pretty.T
@@ -67,6 +68,8 @@
   val brackify_infix: int * lrx -> fixity -> Pretty.T * Pretty.T * Pretty.T -> Pretty.T
   val brackify_block: fixity -> Pretty.T -> Pretty.T list -> Pretty.T -> Pretty.T
   val applify: string -> string -> ('a -> Pretty.T) -> fixity -> Pretty.T -> 'a list -> Pretty.T
+  val tuplify: (fixity -> 'a -> Pretty.T) -> fixity -> 'a list -> Pretty.T option
+
 
   type tyco_syntax
   type simple_const_syntax
@@ -112,6 +115,7 @@
 fun xs @| y = xs @ [y];
 val str = Print_Mode.setmp [] Pretty.str;
 val concat = Pretty.block o Pretty.breaks;
+val commas = Print_Mode.setmp [] Pretty.commas;
 fun enclose l r = Print_Mode.setmp [] (Pretty.enclose l r);
 val brackets = enclose "(" ")" o Pretty.breaks;
 fun enum sep l r = Print_Mode.setmp [] (Pretty.enum sep l r);
@@ -122,7 +126,7 @@
 fun indent i = Print_Mode.setmp [] (Pretty.indent i);
 
 fun string_of_pretty width p = Print_Mode.setmp [] (Pretty.string_of_margin width) p ^ "\n";
-fun writeln_pretty width p = writeln (Print_Mode.setmp [] (Pretty.string_of_margin width) p);
+fun writeln_pretty width p = writeln (string_of_pretty width p);
 
 
 (** names and variable name contexts **)
@@ -242,6 +246,10 @@
       (if (fixity BR fxy_ctxt) then enclose "(" ")" else Pretty.block)
         (p @@ enum "," opn cls (map f ps));
 
+fun tuplify _ _ [] = NONE
+  | tuplify print fxy [x] = SOME (print fxy x)
+  | tuplify print _ xs = SOME (enum "," "(" ")" (map (print NOBR) xs));
+
 
 (* generic syntax *)
 
@@ -276,8 +284,8 @@
       fold_map (Code_Thingol.ensure_declared_const thy) cs naming
       |-> (fn cs' => pair (Complex_const_syntax (n, f literals cs')));
 
-fun gen_print_app print_app_expr print_term syntax_const some_thm vars fxy (app as ((c, (_, function_typs)), ts)) =
-  case syntax_const c
+fun gen_print_app print_app_expr print_term const_syntax some_thm vars fxy (app as ((c, (_, function_typs)), ts)) =
+  case const_syntax c
    of NONE => brackify fxy (print_app_expr some_thm vars app)
     | SOME (Plain_const_syntax (_, s)) => brackify fxy (str s :: map (print_term some_thm vars BR) ts)
     | SOME (Complex_const_syntax (k, print)) =>
--- a/src/Tools/Code/code_scala.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/Code/code_scala.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -24,15 +24,14 @@
 
 (** Scala serializer **)
 
-fun print_scala_stmt labelled_name syntax_tyco syntax_const reserved
-    args_num is_singleton_constr deresolve =
+fun print_scala_stmt labelled_name tyco_syntax const_syntax reserved
+    args_num is_singleton_constr (deresolve, deresolve_full) =
   let
-    val deresolve_base = Long_Name.base_name o deresolve;
     fun lookup_tyvar tyvars = lookup_var tyvars o first_upper;
     fun intro_tyvars vs = intro_vars (map (first_upper o fst) vs);
     fun print_tyco_expr tyvars fxy (tyco, tys) = applify "[" "]"
           (print_typ tyvars NOBR) fxy ((str o deresolve) tyco) tys
-    and print_typ tyvars fxy (tyco `%% tys) = (case syntax_tyco tyco
+    and print_typ tyvars fxy (tyco `%% tys) = (case tyco_syntax tyco
          of NONE => print_tyco_expr tyvars fxy (tyco, tys)
           | SOME (i, print) => print (print_typ tyvars) fxy tys)
       | print_typ tyvars fxy (ITyVar v) = (str o lookup_tyvar tyvars) v;
@@ -68,7 +67,7 @@
           end 
       | print_term tyvars is_pat some_thm vars fxy (ICase (cases as (_, t0))) =
           (case Code_Thingol.unfold_const_app t0
-           of SOME (c_ts as ((c, _), _)) => if is_none (syntax_const c)
+           of SOME (c_ts as ((c, _), _)) => if is_none (const_syntax c)
                 then print_case tyvars some_thm vars fxy cases
                 else print_app tyvars is_pat some_thm vars fxy c_ts
             | NONE => print_case tyvars some_thm vars fxy cases)
@@ -77,8 +76,8 @@
       let
         val k = length ts;
         val arg_typs' = if is_pat orelse
-          (is_none (syntax_const c) andalso is_singleton_constr c) then [] else arg_typs;
-        val (l, print') = case syntax_const c
+          (is_none (const_syntax c) andalso is_singleton_constr c) then [] else arg_typs;
+        val (l, print') = case const_syntax c
          of NONE => (args_num c, fn fxy => fn ts => applify "(" ")"
               (print_term tyvars is_pat some_thm vars NOBR) fxy
                 (applify "[" "]" (print_typ tyvars NOBR)
@@ -157,7 +156,7 @@
               fold Code_Thingol.add_tyconames (t :: ts)) eqs [];
             val tyvars = reserved
               |> intro_base_names
-                   (is_none o syntax_tyco) deresolve tycos
+                   (is_none o tyco_syntax) deresolve tycos
               |> intro_tyvars vs;
             val simple = case eqs
              of [((ts, _), _)] => forall Code_Thingol.is_IVar ts
@@ -166,14 +165,14 @@
               (map (snd o fst) eqs) [];
             val vars1 = reserved
               |> intro_base_names
-                   (is_none o syntax_const) deresolve consts
+                   (is_none o const_syntax) deresolve consts
             val params = if simple
               then (map (fn IVar (SOME x) => x) o fst o fst o hd) eqs
               else aux_params vars1 (map (fst o fst) eqs);
             val vars2 = intro_vars params vars1;
             val (tys', ty') = Code_Thingol.unfold_fun_n (length params) ty;
-            fun print_tuple [p] = p
-              | print_tuple ps = enum "," "(" ")" ps;
+            fun tuplify [p] = p
+              | tuplify ps = enum "," "(" ")" ps;
             fun print_rhs vars' ((_, t), (some_thm, _)) =
               print_term tyvars false some_thm vars' NOBR t;
             fun print_clause (eq as ((ts, _), (some_thm, _))) =
@@ -182,7 +181,7 @@
                   (insert (op =)) ts []) vars1;
               in
                 concat [str "case",
-                  print_tuple (map (print_term tyvars true some_thm vars' NOBR) ts),
+                  tuplify (map (print_term tyvars true some_thm vars' NOBR) ts),
                   str "=>", print_rhs vars' eq]
               end;
             val head = print_defhead tyvars vars2 name vs params tys' ty';
@@ -190,33 +189,34 @@
             concat [head, print_rhs vars2 (hd eqs)]
           else
             Pretty.block_enclose
-              (concat [head, print_tuple (map (str o lookup_var vars2) params),
+              (concat [head, tuplify (map (str o lookup_var vars2) params),
                 str "match", str "{"], str "}")
               (map print_clause eqs)
           end;
-    val print_method = (str o Library.enclose "`" "+`" o deresolve_base);
+    val print_method = str o Library.enclose "`" "`" o space_implode "+"
+      o Long_Name.explode o deresolve_full;
     fun print_stmt (name, Code_Thingol.Fun (_, (((vs, ty), raw_eqs), _))) =
           print_def name (vs, ty) (filter (snd o snd) raw_eqs)
       | print_stmt (name, Code_Thingol.Datatype (_, (vs, cos))) =
           let
             val tyvars = intro_tyvars vs reserved;
             fun print_co ((co, _), []) =
-                  concat [str "final", str "case", str "object", (str o deresolve_base) co,
-                    str "extends", applify "[" "]" I NOBR ((str o deresolve_base) name)
+                  concat [str "final", str "case", str "object", (str o deresolve) co,
+                    str "extends", applify "[" "]" I NOBR ((str o deresolve) name)
                       (replicate (length vs) (str "Nothing"))]
               | print_co ((co, vs_args), tys) =
                   concat [applify "(" ")"
                     (fn (v, arg) => constraint (str v) (print_typ tyvars NOBR arg)) NOBR
                     (applify "[" "]" (str o lookup_tyvar tyvars) NOBR ((concat o map str)
-                      ["final", "case", "class", deresolve_base co]) vs_args)
+                      ["final", "case", "class", deresolve co]) vs_args)
                     (Name.names (snd reserved) "a" tys),
                     str "extends",
                     applify "[" "]" (str o lookup_tyvar tyvars o fst) NOBR
-                      ((str o deresolve_base) name) vs
+                      ((str o deresolve) name) vs
                   ];
           in
             Pretty.chunks (applify "[" "]" (str o prefix "+" o lookup_tyvar tyvars o fst)
-              NOBR ((concat o map str) ["abstract", "sealed", "class", deresolve_base name]) vs
+              NOBR ((concat o map str) ["abstract", "sealed", "class", deresolve name]) vs
                 :: map print_co cos)
           end
       | print_stmt (name, Code_Thingol.Class (_, (v, (super_classes, classparams)))) =
@@ -249,7 +249,7 @@
           in
             Pretty.chunks (
               (Pretty.block_enclose
-                (concat ([str "trait", (add_typarg o deresolve_base) name]
+                (concat ([str "trait", (add_typarg o deresolve) name]
                   @ the_list (print_super_classes super_classes) @ [str "{"]), str "}")
                 (map print_classparam_val classparams))
               :: map print_classparam_def classparams
@@ -281,67 +281,60 @@
           end;
   in print_stmt end;
 
-fun scala_program_of_program labelled_name module_name reserved raw_module_alias program =
+fun scala_program_of_program labelled_name reserved module_alias program =
   let
-    val the_module_name = the_default "Program" module_name;
-    val module_alias = K (SOME the_module_name);
-    val reserved = Name.make_context reserved;
-    fun prepare_stmt (name, stmt) (nsps, stmts) =
+    fun namify_module name_fragment ((nsp_class, nsp_object), nsp_common) =
+      let
+        val declare = Name.declare name_fragment;
+      in (name_fragment, ((declare nsp_class, declare nsp_object), declare nsp_common)) end;
+    fun namify_class base ((nsp_class, nsp_object), nsp_common) =
+      let
+        val (base', nsp_class') = yield_singleton Name.variants base nsp_class
+      in (base', ((nsp_class', nsp_object), Name.declare base' nsp_common)) end;
+    fun namify_object base ((nsp_class, nsp_object), nsp_common) =
+      let
+        val (base', nsp_object') = yield_singleton Name.variants base nsp_object
+      in (base', ((nsp_class, nsp_object'), Name.declare base' nsp_common)) end;
+    fun namify_common upper base ((nsp_class, nsp_object), nsp_common) =
       let
-        val (_, base) = Code_Printer.dest_name name;
-        val mk_name_stmt = yield_singleton Name.variants;
-        fun add_class ((nsp_class, nsp_object), nsp_common) =
-          let
-            val (base', nsp_class') = mk_name_stmt base nsp_class
-          in (base', ((nsp_class', nsp_object), Name.declare base' nsp_common)) end;
-        fun add_object ((nsp_class, nsp_object), nsp_common) =
-          let
-            val (base', nsp_object') = mk_name_stmt base nsp_object
-          in (base', ((nsp_class, nsp_object'), Name.declare base' nsp_common)) end;
-        fun add_common upper ((nsp_class, nsp_object), nsp_common) =
-          let
-            val (base', nsp_common') =
-              mk_name_stmt (if upper then first_upper base else base) nsp_common
-          in
-            (base',
-              ((Name.declare base' nsp_class, Name.declare base' nsp_object), nsp_common'))
-          end;
-        val add_name = case stmt
-         of Code_Thingol.Fun _ => add_object
-          | Code_Thingol.Datatype _ => add_class
-          | Code_Thingol.Datatypecons _ => add_common true
-          | Code_Thingol.Class _ => add_class
-          | Code_Thingol.Classrel _ => add_object
-          | Code_Thingol.Classparam _ => add_object
-          | Code_Thingol.Classinst _ => add_common false;
-        fun add_stmt base' = case stmt
-         of Code_Thingol.Datatypecons _ => cons (name, (base', NONE))
-          | Code_Thingol.Classrel _ => cons (name, (base', NONE))
-          | Code_Thingol.Classparam _ => cons (name, (base', NONE))
-          | _ => cons (name, (base', SOME stmt));
+        val (base', nsp_common') =
+          yield_singleton Name.variants (if upper then first_upper base else base) nsp_common
       in
-        nsps
-        |> add_name
-        |-> (fn base' => rpair (add_stmt base' stmts))
+        (base',
+          ((Name.declare base' nsp_class, Name.declare base' nsp_object), nsp_common'))
       end;
-    val stmts = AList.make (Graph.get_node program) (Graph.strong_conn program |> flat)
-      |> filter_out (Code_Thingol.is_case o snd);
-    val (_, sca_program) = fold prepare_stmt stmts (((reserved, reserved), reserved), []);
-    fun deresolver name = (fst o the o AList.lookup (op =) sca_program) name
-      handle Option => error ("Unknown statement name: " ^ labelled_name name);
-  in (deresolver, (the_module_name, sca_program)) end;
+    fun namify_stmt (Code_Thingol.Fun _) = namify_object
+      | namify_stmt (Code_Thingol.Datatype _) = namify_class
+      | namify_stmt (Code_Thingol.Datatypecons _) = namify_common true
+      | namify_stmt (Code_Thingol.Class _) = namify_class
+      | namify_stmt (Code_Thingol.Classrel _) = namify_object
+      | namify_stmt (Code_Thingol.Classparam _) = namify_object
+      | namify_stmt (Code_Thingol.Classinst _) = namify_common false;
+    fun memorize_implicits name =
+      let
+        fun is_classinst stmt = case stmt
+         of Code_Thingol.Classinst _ => true
+          | _ => false;
+        val implicits = filter (is_classinst o Graph.get_node program)
+          (Graph.imm_succs program name);
+      in union (op =) implicits end;
+  in
+    Code_Namespace.hierarchical_program labelled_name { module_alias = module_alias, reserved = reserved,
+      empty_nsp = ((reserved, reserved), reserved), namify_module = namify_module, namify_stmt = namify_stmt,
+      cyclic_modules = true, empty_data = [], memorize_data = memorize_implicits } program
+  end;
 
-fun serialize_scala raw_module_name labelled_name
-    raw_reserved includes raw_module_alias
-    _ syntax_tyco syntax_const (code_of_pretty, code_writeln)
-    program stmt_names destination =
+fun serialize_scala { labelled_name, reserved_syms, includes,
+    module_alias, class_syntax, tyco_syntax, const_syntax, program,
+    names, presentation_names } =
   let
-    val presentation_stmt_names = Code_Target.stmt_names_of_destination destination;
-    val module_name = if null presentation_stmt_names then raw_module_name else SOME "Code";
-    val reserved = fold (insert (op =) o fst) includes raw_reserved;
-    val (deresolver, (the_module_name, sca_program)) = scala_program_of_program labelled_name
-      module_name reserved raw_module_alias program;
-    val reserved = make_vars reserved;
+
+    (* build program *)
+    val reserved = fold (insert (op =) o fst) includes reserved_syms;
+    val { deresolver, hierarchical_program = sca_program } =
+      scala_program_of_program labelled_name (Name.make_context reserved) module_alias program;
+
+    (* print statements *)
     fun lookup_constr tyco constr = case Graph.get_node program tyco
      of Code_Thingol.Datatype (_, (_, constrs)) =>
           the (AList.lookup (op = o apsnd fst) constrs constr);
@@ -358,45 +351,54 @@
     fun is_singleton_constr c = case Graph.get_node program c
      of Code_Thingol.Datatypecons (_, tyco) => null (lookup_constr tyco c)
       | _ => false;
-    val print_stmt = print_scala_stmt labelled_name syntax_tyco syntax_const
-      reserved args_num is_singleton_constr deresolver;
-    fun print_module name imports content =
-      (name, Pretty.chunks (
-        str ("object " ^ name ^ " {")
-        :: (if null imports then []
-          else str "" :: map (fn name => str ("import " ^ name ^ "._")) imports)
-        @ [str "",
-        content,
-        str "",
-        str "}"]
-      ));
-    fun serialize_module the_module_name sca_program =
+    val print_stmt = print_scala_stmt labelled_name tyco_syntax const_syntax
+      (make_vars reserved) args_num is_singleton_constr;
+
+    (* print nodes *)
+    fun print_module base implicit_ps p = Pretty.chunks2
+      ([str ("object " ^ base ^ " {")]
+        @ (if null implicit_ps then [] else (single o Pretty.block)
+            (str "import /*implicits*/" :: Pretty.brk 1 :: commas implicit_ps))
+        @ [p, str ("} /* object " ^ base ^ " */")]);
+    fun print_implicit prefix_fragments implicit =
       let
-        val content = Pretty.chunks2 (map_filter
-          (fn (name, (_, SOME stmt)) => SOME (print_stmt (name, stmt))
-            | (_, (_, NONE)) => NONE) sca_program);
-      in print_module the_module_name (map fst includes) content end;
-    fun check_destination destination =
-      (File.check destination; destination);
-    fun write_module destination (modlname, content) =
-      let
-        val filename = case modlname
-         of "" => Path.explode "Main.scala"
-          | _ => (Path.ext "scala" o Path.explode o implode o separate "/"
-                o Long_Name.explode) modlname;
-        val pathname = Path.append destination filename;
-        val _ = File.mkdir_leaf (Path.dir pathname);
-      in File.write pathname (code_of_pretty content) end
+        val s = deresolver prefix_fragments implicit;
+      in if length (Long_Name.explode s) = 1 then NONE else SOME (str s) end;
+    fun print_node _ (_, Dummy) = NONE
+      | print_node prefix_fragments (name, Code_Namespace.Stmt stmt) =
+          if null presentation_names
+          orelse member (op =) presentation_names name
+          then SOME (print_stmt (deresolver prefix_fragments, deresolver []) (name, stmt))
+          else NONE
+      | print_node prefix_fragments (name_fragment, Code_Namespace.Module (implicits, nodes)) =
+          if null presentation_names
+          then
+            let
+              val prefix_fragments' = prefix_fragments @ [name_fragment];
+            in
+              Option.map (print_module name_fragment
+                (map_filter (print_implicit prefix_fragments') implicits))
+                  (print_nodes prefix_fragments' nodes)
+            end
+          else print_nodes [] nodes
+    and print_nodes prefix_fragments nodes = let
+        val ps = map_filter (fn name => print_node prefix_fragments (name,
+          snd (Graph.get_node nodes name)))
+            ((rev o flat o Graph.strong_conn) nodes);
+      in if null ps then NONE else SOME (Pretty.chunks2 ps) end;
+
+    (* serialization *)
+    val p_includes = if null presentation_names then map snd includes else [];
+    val p = Pretty.chunks2 (p_includes @ the_list (print_nodes [] sca_program));
+    fun write width NONE = writeln_pretty width
+      | write width (SOME p) = File.write p o string_of_pretty width;
   in
-    Code_Target.mk_serialization target
-      (fn NONE => K () o map (code_writeln o snd) | SOME file => K () o map
-        (write_module (check_destination file)))
-      (rpair [] o cat_lines o map (code_of_pretty o snd))
-      (map (fn (name, content) => print_module name [] content) includes
-        @| serialize_module the_module_name sca_program)
-      destination
+    Code_Target.serialization write (rpair [] oo string_of_pretty) p
   end;
 
+val serializer : Code_Target.serializer =
+  Code_Target.parse_args (Scan.succeed ()) #> K serialize_scala;
+
 val literals = let
   fun char_scala c = if c = "'" then "\\'"
     else if c = "\"" then "\\\""
@@ -422,18 +424,14 @@
 
 (** Isar setup **)
 
-fun isar_serializer module_name =
-  Code_Target.parse_args (Scan.succeed ())
-  #> (fn () => serialize_scala module_name);
-
 val setup =
   Code_Target.add_target
-    (target, { serializer = isar_serializer, literals = literals,
-      check = { env_var = "SCALA_HOME", make_destination = I,
-        make_command = fn scala_home => fn p => fn _ =>
+    (target, { serializer = serializer, literals = literals,
+      check = { env_var = "SCALA_HOME", make_destination = fn p => Path.append p (Path.explode "ROOT.scala"),
+        make_command = fn scala_home => fn _ =>
           "export JAVA_OPTS='-Xms128m -Xmx512m -Xss2m' && "
-            ^ Path.implode (Path.append (Path.explode scala_home) (Path.explode "bin/scalac")) ^ " *.scala" } })
-  #> Code_Target.add_syntax_tyco target "fun"
+            ^ Path.implode (Path.append (Path.explode scala_home) (Path.explode "bin/scalac")) ^ " ROOT.scala" } })
+  #> Code_Target.add_tyco_syntax target "fun"
      (SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>
         brackify_infix (1, R) fxy (
           print_typ BR ty1 (*product type vs. tupled arguments!*),
--- a/src/Tools/Code/code_simp.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/Code/code_simp.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -37,7 +37,8 @@
 
 (* dedicated simpset *)
 
-structure Simpset = Theory_Data (
+structure Simpset = Theory_Data
+(
   type T = simpset;
   val empty = empty_ss;
   fun extend ss = MetaSimplifier.inherit_context empty_ss ss;
--- a/src/Tools/Code/code_target.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/Code/code_target.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -1,7 +1,7 @@
 (*  Title:      Tools/Code/code_target.ML
     Author:     Florian Haftmann, TU Muenchen
 
-Serializer from intermediate language ("Thin-gol") to target languages.
+Generic infrastructure for target language data.
 *)
 
 signature CODE_TARGET =
@@ -9,43 +9,50 @@
   val cert_tyco: theory -> string -> string
   val read_tyco: theory -> string -> string
 
+  val export_code_for: theory -> Path.T option -> string -> int option -> string -> Token.T list
+    -> Code_Thingol.naming -> Code_Thingol.program -> string list -> unit
+  val produce_code_for: theory -> string -> int option -> string -> Token.T list
+    -> Code_Thingol.naming -> Code_Thingol.program -> string list -> string * string option list
+  val present_code_for: theory -> string -> int option -> string -> Token.T list
+    -> Code_Thingol.naming -> Code_Thingol.program -> string list * string list -> string
+  val check_code_for: theory -> string -> bool -> Token.T list
+    -> Code_Thingol.naming -> Code_Thingol.program -> string list -> unit
+
+  val export_code: theory -> string list
+    -> (((string * string) * Path.T option) * Token.T list) list -> unit
+  val produce_code: theory -> string list
+    -> string -> int option -> string -> Token.T list -> string * string option list
+  val present_code: theory -> string list -> (Code_Thingol.naming -> string list)
+    -> string -> int option -> string -> Token.T list -> string
+  val check_code: theory -> string list
+    -> ((string * bool) * Token.T list) list -> unit
+
+  val shell_command: string (*theory name*) -> string (*export_code expr*) -> unit
+
   type serializer
   type literals = Code_Printer.literals
   val add_target: string * { serializer: serializer, literals: literals,
     check: { env_var: string, make_destination: Path.T -> Path.T,
-      make_command: string -> Path.T -> string -> string } } -> theory -> theory
+      make_command: string -> string -> string } } -> theory -> theory
   val extend_target: string *
       (string * (Code_Thingol.naming -> Code_Thingol.program -> Code_Thingol.program))
     -> theory -> theory
   val assert_target: theory -> string -> string
-
-  type destination
+  val the_literals: theory -> string -> literals
   type serialization
   val parse_args: 'a parser -> Token.T list -> 'a
-  val stmt_names_of_destination: destination -> string list
-  val mk_serialization: string -> (Path.T option -> 'a -> unit)
-    -> ('a -> string * string option list)
+  val serialization: (int -> Path.T option -> 'a -> unit)
+    -> (int -> 'a -> string * string option list)
     -> 'a -> serialization
-  val serialize: theory -> string -> int option -> string option -> Token.T list
-    -> Code_Thingol.naming -> Code_Thingol.program -> string list -> serialization
-  val serialize_custom: theory -> string * serializer
-    -> Code_Thingol.naming -> Code_Thingol.program -> string list -> string * string option list
-  val the_literals: theory -> string -> literals
-  val export: serialization -> unit
-  val file: Path.T -> serialization -> unit
-  val string: string list -> serialization -> string
-  val code_of: theory -> string -> int option -> string
-    -> string list -> (Code_Thingol.naming -> string list) -> string
   val set_default_code_width: int -> theory -> theory
-  val shell_command: string (*theory name*) -> string (*export_code expr*) -> unit
 
   val allow_abort: string -> theory -> theory
   type tyco_syntax = Code_Printer.tyco_syntax
   type const_syntax = Code_Printer.const_syntax
-  val add_syntax_class: string -> class -> string option -> theory -> theory
-  val add_syntax_inst: string -> class * string -> unit option -> theory -> theory
-  val add_syntax_tyco: string -> string -> tyco_syntax option -> theory -> theory
-  val add_syntax_const: string -> string -> const_syntax option -> theory -> theory
+  val add_class_syntax: string -> class -> string option -> theory -> theory
+  val add_instance_syntax: string -> class * string -> unit option -> theory -> theory
+  val add_tyco_syntax: string -> string -> tyco_syntax option -> theory -> theory
+  val add_const_syntax: string -> string -> const_syntax option -> theory -> theory
   val add_reserved: string -> string -> theory -> theory
   val add_include: string -> string * (string * string list) option -> theory -> theory
 end;
@@ -60,90 +67,91 @@
 type const_syntax = Code_Printer.const_syntax;
 
 
-(** basics **)
-
-datatype destination = Export | File of Path.T | String of string list;
-type serialization = destination -> (string * string option list) option;
+(** abstract nonsense **)
 
-fun export f = (f Export; ());
-fun file p f = (f (File p); ());
-fun string stmts f = fst (the (f (String stmts)));
+datatype destination = Export of Path.T option | Produce | Present of string list;
+type serialization = int -> destination -> (string * string option list) option;
 
-fun stmt_names_of_destination (String stmts) = stmts
+fun stmt_names_of_destination (Present stmt_names) = stmt_names
   | stmt_names_of_destination _ = [];
 
-fun mk_serialization target output _ code Export = (output NONE code ; NONE)
-  | mk_serialization target output _ code (File file) = (output (SOME file) code; NONE)
-  | mk_serialization target _ string code (String _) = SOME (string code);
+fun serialization output _ content width (Export some_path) = (output width some_path content; NONE)
+  | serialization _ string content width Produce = SOME (string width content)
+  | serialization _ string content width (Present _) = SOME (string width content);
+
+fun export some_path f = (f (Export some_path); ());
+fun produce f = the (f Produce);
+fun present stmt_names f = fst (the (f (Present stmt_names)));
 
 
 (** theory data **)
 
-datatype name_syntax_table = NameSyntaxTable of {
+datatype symbol_syntax_data = Symbol_Syntax_Data of {
   class: string Symtab.table,
   instance: unit Symreltab.table,
   tyco: Code_Printer.tyco_syntax Symtab.table,
   const: Code_Printer.const_syntax Symtab.table
 };
 
-fun mk_name_syntax_table ((class, instance), (tyco, const)) =
-  NameSyntaxTable { class = class, instance = instance, tyco = tyco, const = const };
-fun map_name_syntax_table f (NameSyntaxTable { class, instance, tyco, const }) =
-  mk_name_syntax_table (f ((class, instance), (tyco, const)));
-fun merge_name_syntax_table (NameSyntaxTable { class = class1, instance = instance1, tyco = tyco1, const = const1 },
-    NameSyntaxTable { class = class2, instance = instance2, tyco = tyco2, const = const2 }) =
-  mk_name_syntax_table (
+fun make_symbol_syntax_data ((class, instance), (tyco, const)) =
+  Symbol_Syntax_Data { class = class, instance = instance, tyco = tyco, const = const };
+fun map_symbol_syntax_data f (Symbol_Syntax_Data { class, instance, tyco, const }) =
+  make_symbol_syntax_data (f ((class, instance), (tyco, const)));
+fun merge_symbol_syntax_data
+  (Symbol_Syntax_Data { class = class1, instance = instance1, tyco = tyco1, const = const1 },
+    Symbol_Syntax_Data { class = class2, instance = instance2, tyco = tyco2, const = const2 }) =
+  make_symbol_syntax_data (
     (Symtab.join (K snd) (class1, class2),
        Symreltab.join (K snd) (instance1, instance2)),
     (Symtab.join (K snd) (tyco1, tyco2),
        Symtab.join (K snd) (const1, const2))
   );
 
-type serializer =
-  string option                         (*module name*)
-  -> Token.T list                       (*arguments*)
-  -> (string -> string)                 (*labelled_name*)
-  -> string list                        (*reserved symbols*)
-  -> (string * Pretty.T) list           (*includes*)
-  -> (string -> string option)          (*module aliasses*)
-  -> (string -> string option)          (*class syntax*)
-  -> (string -> Code_Printer.tyco_syntax option)
-  -> (string -> Code_Printer.activated_const_syntax option)
-  -> ((Pretty.T -> string) * (Pretty.T -> unit))
-  -> Code_Thingol.program
-  -> string list                        (*selected statements*)
+type serializer = Token.T list (*arguments*) -> {
+    labelled_name: string -> string,
+    reserved_syms: string list,
+    includes: (string * Pretty.T) list,
+    module_alias: string -> string option,
+    class_syntax: string -> string option,
+    tyco_syntax: string -> Code_Printer.tyco_syntax option,
+    const_syntax: string -> Code_Printer.activated_const_syntax option,
+    program: Code_Thingol.program,
+    names: string list,
+    presentation_names: string list }
   -> serialization;
 
-datatype description = Fundamental of { serializer: serializer, literals: Code_Printer.literals,
+datatype description = Fundamental of { serializer: serializer,
+      literals: literals,
       check: { env_var: string, make_destination: Path.T -> Path.T,
-        make_command: string -> Path.T -> string -> string } }
-  | Extension of string * (Code_Thingol.naming -> Code_Thingol.program -> Code_Thingol.program);
+        make_command: string -> string -> string } }
+  | Extension of string *
+      (Code_Thingol.naming -> Code_Thingol.program -> Code_Thingol.program);
 
 datatype target = Target of {
   serial: serial,
   description: description,
   reserved: string list,
   includes: (Pretty.T * string list) Symtab.table,
-  name_syntax_table: name_syntax_table,
-  module_alias: string Symtab.table
+  module_alias: string Symtab.table,
+  symbol_syntax: symbol_syntax_data
 };
 
-fun make_target ((serial, description), ((reserved, includes), (name_syntax_table, module_alias))) =
+fun make_target ((serial, description), ((reserved, includes), (module_alias, symbol_syntax))) =
   Target { serial = serial, description = description, reserved = reserved, 
-    includes = includes, name_syntax_table = name_syntax_table, module_alias = module_alias };
-fun map_target f ( Target { serial, description, reserved, includes, name_syntax_table, module_alias } ) =
-  make_target (f ((serial, description), ((reserved, includes), (name_syntax_table, module_alias))));
+    includes = includes, module_alias = module_alias, symbol_syntax = symbol_syntax };
+fun map_target f ( Target { serial, description, reserved, includes, module_alias, symbol_syntax } ) =
+  make_target (f ((serial, description), ((reserved, includes), (module_alias, symbol_syntax))));
 fun merge_target strict target (Target { serial = serial1, description = description,
   reserved = reserved1, includes = includes1,
-  name_syntax_table = name_syntax_table1, module_alias = module_alias1 },
+  module_alias = module_alias1, symbol_syntax = symbol_syntax1 },
     Target { serial = serial2, description = _,
       reserved = reserved2, includes = includes2,
-      name_syntax_table = name_syntax_table2, module_alias = module_alias2 }) =
+      module_alias = module_alias2, symbol_syntax = symbol_syntax2 }) =
   if serial1 = serial2 orelse not strict then
     make_target ((serial1, description),
       ((merge (op =) (reserved1, reserved2), Symtab.join (K snd) (includes1, includes2)),
-        (merge_name_syntax_table (name_syntax_table1, name_syntax_table2),
-          Symtab.join (K snd) (module_alias1, module_alias2))
+        (Symtab.join (K snd) (module_alias1, module_alias2),
+          merge_symbol_syntax_data (symbol_syntax1, symbol_syntax2))
     ))
   else
     error ("Incompatible targets: " ^ quote target);
@@ -151,8 +159,8 @@
 fun the_description (Target { description, ... }) = description;
 fun the_reserved (Target { reserved, ... }) = reserved;
 fun the_includes (Target { includes, ... }) = includes;
-fun the_name_syntax (Target { name_syntax_table = NameSyntaxTable x, ... }) = x;
 fun the_module_alias (Target { module_alias , ... }) = module_alias;
+fun the_symbol_syntax (Target { symbol_syntax = Symbol_Syntax_Data x, ... }) = x;
 
 structure Targets = Theory_Data
 (
@@ -190,8 +198,8 @@
     thy
     |> (Targets.map o apfst o apfst o Symtab.update)
           (target, make_target ((serial (), seri), (([], Symtab.empty),
-            (mk_name_syntax_table ((Symtab.empty, Symreltab.empty), (Symtab.empty, Symtab.empty)),
-              Symtab.empty))))
+            (Symtab.empty, make_symbol_syntax_data ((Symtab.empty, Symreltab.empty),
+              (Symtab.empty, Symtab.empty))))))
   end;
 
 fun add_target (target, seri) = put_target (target, Fundamental seri);
@@ -210,10 +218,10 @@
   map_target_data target o apsnd o apfst o apfst;
 fun map_includes target =
   map_target_data target o apsnd o apfst o apsnd;
-fun map_name_syntax target =
-  map_target_data target o apsnd o apsnd o apfst o map_name_syntax_table;
 fun map_module_alias target =
-  map_target_data target o apsnd o apsnd o apsnd;
+  map_target_data target o apsnd o apsnd o apfst;
+fun map_symbol_syntax target =
+  map_target_data target o apsnd o apsnd o apsnd o map_symbol_syntax_data;
 
 fun set_default_code_width k = (Targets.map o apsnd) (K k);
 
@@ -236,6 +244,23 @@
 
 local
 
+fun activate_target_for thy target naming program =
+  let
+    val ((targets, abortable), default_width) = Targets.get thy;
+    fun collapse_hierarchy target =
+      let
+        val data = case Symtab.lookup targets target
+         of SOME data => data
+          | NONE => error ("Unknown code target language: " ^ quote target);
+      in case the_description data
+       of Fundamental _ => (I, data)
+        | Extension (super, modify) => let
+            val (modify', data') = collapse_hierarchy super
+          in (modify' #> modify naming, merge_target false target (data', data)) end
+      end;
+    val (modify, data) = collapse_hierarchy target;
+  in (default_width, abortable, data, modify program) end;
+
 fun activate_syntax lookup_name src_tab = Symtab.empty
   |> fold_map (fn thing_identifier => fn tab => case lookup_name thing_identifier
        of SOME name => (SOME name,
@@ -253,52 +278,66 @@
         | NONE => (NONE, (tab, naming))) (Symtab.keys src_tab)
   |>> map_filter I;
 
-fun invoke_serializer thy abortable serializer literals reserved abs_includes 
-    module_alias class instance tyco const module width args naming program2 names1 =
+fun activate_symbol_syntax thy literals naming
+    class_syntax instance_syntax tyco_syntax const_syntax =
   let
-    val (names_class, class') =
-      activate_syntax (Code_Thingol.lookup_class naming) class;
+    val (names_class, class_syntax') =
+      activate_syntax (Code_Thingol.lookup_class naming) class_syntax;
     val names_inst = map_filter (Code_Thingol.lookup_instance naming)
-      (Symreltab.keys instance);
-    val (names_tyco, tyco') =
-      activate_syntax (Code_Thingol.lookup_tyco naming) tyco;
-    val (names_const, (const', _)) =
-      activate_const_syntax thy literals const naming;
-    val names_hidden = names_class @ names_inst @ names_tyco @ names_const;
+      (Symreltab.keys instance_syntax);
+    val (names_tyco, tyco_syntax') =
+      activate_syntax (Code_Thingol.lookup_tyco naming) tyco_syntax;
+    val (names_const, (const_syntax', _)) =
+      activate_const_syntax thy literals const_syntax naming;
+  in
+    (names_class @ names_inst @ names_tyco @ names_const,
+      (class_syntax', tyco_syntax', const_syntax'))
+  end;
+
+fun project_program thy abortable names_hidden names1 program2 =
+  let
     val names2 = subtract (op =) names_hidden names1;
     val program3 = Graph.subgraph (not o member (op =) names_hidden) program2;
-    val names_all = Graph.all_succs program3 names2;
-    val includes = abs_includes names_all;
-    val program4 = Graph.subgraph (member (op =) names_all) program3;
+    val names4 = Graph.all_succs program3 names2;
     val empty_funs = filter_out (member (op =) abortable)
       (Code_Thingol.empty_funs program3);
     val _ = if null empty_funs then () else error ("No code equations for "
       ^ commas (map (Sign.extern_const thy) empty_funs));
+    val program4 = Graph.subgraph (member (op =) names4) program3;
+  in (names4, program4) end;
+
+fun invoke_serializer thy abortable serializer literals reserved abs_includes 
+    module_alias proto_class_syntax proto_instance_syntax proto_tyco_syntax proto_const_syntax
+    module_name args naming proto_program (names, presentation_names) =
+  let
+    val (names_hidden, (class_syntax, tyco_syntax, const_syntax)) =
+      activate_symbol_syntax thy literals naming
+        proto_class_syntax proto_instance_syntax proto_tyco_syntax proto_const_syntax;
+    val (names_all, program) = project_program thy abortable names_hidden names proto_program;
+    val includes = abs_includes names_all;
   in
-    serializer module args (Code_Thingol.labelled_name thy program2) reserved includes
-      (Symtab.lookup module_alias) (Symtab.lookup class')
-      (Symtab.lookup tyco') (Symtab.lookup const')
-      (Code_Printer.string_of_pretty width, Code_Printer.writeln_pretty width)
-      program4 names1
+    serializer args {
+      labelled_name = Code_Thingol.labelled_name thy proto_program,
+      reserved_syms = reserved,
+      includes = includes,
+      module_alias = if module_name = "" then Symtab.lookup module_alias else K (SOME module_name),
+      class_syntax = Symtab.lookup class_syntax,
+      tyco_syntax = Symtab.lookup tyco_syntax,
+      const_syntax = Symtab.lookup const_syntax,
+      program = program,
+      names = names,
+      presentation_names = presentation_names }
   end;
 
-fun mount_serializer thy alt_serializer target some_width module args naming program names =
+fun mount_serializer thy target some_width raw_module_name args naming proto_program names destination =
   let
-    val ((targets, abortable), default_width) = Targets.get thy;
-    fun collapse_hierarchy target =
-      let
-        val data = case Symtab.lookup targets target
-         of SOME data => data
-          | NONE => error ("Unknown code target language: " ^ quote target);
-      in case the_description data
-       of Fundamental _ => (I, data)
-        | Extension (super, modify) => let
-            val (modify', data') = collapse_hierarchy super
-          in (modify' #> modify naming, merge_target false target (data', data)) end
-      end;
-    val (modify, data) = collapse_hierarchy target;
-    val serializer = the_default (case the_description data
-     of Fundamental seri => #serializer seri) alt_serializer;
+    val (default_width, abortable, data, program) =
+      activate_target_for thy target naming proto_program;
+    val serializer = case the_description data
+     of Fundamental seri => #serializer seri;
+    val presentation_names = stmt_names_of_destination destination;
+    val module_name = if null presentation_names
+      then raw_module_name else "Code";
     val reserved = the_reserved data;
     fun select_include names_all (name, (content, cs)) =
       if null cs then SOME (name, content)
@@ -308,31 +347,42 @@
       then SOME (name, content) else NONE;
     fun includes names_all = map_filter (select_include names_all)
       ((Symtab.dest o the_includes) data);
-    val module_alias = the_module_alias data;
-    val { class, instance, tyco, const } = the_name_syntax data;
+    val module_alias = the_module_alias data 
+    val { class, instance, tyco, const } = the_symbol_syntax data;
     val literals = the_literals thy target;
     val width = the_default default_width some_width;
   in
     invoke_serializer thy abortable serializer literals reserved
-      includes module_alias class instance tyco const module width args naming (modify program) names
+      includes module_alias class instance tyco const module_name args
+        naming program (names, presentation_names) width destination
   end;
 
+fun assert_module_name "" = error ("Empty module name not allowed.")
+  | assert_module_name module_name = module_name;
+
 in
 
-fun serialize thy = mount_serializer thy NONE;
+fun export_code_for thy some_path target some_width some_module_name args naming program names =
+  export some_path (mount_serializer thy target some_width some_module_name args naming program names);
+
+fun produce_code_for thy target some_width module_name args naming program names =
+  produce (mount_serializer thy target some_width (assert_module_name module_name) args naming program names);
 
-fun check thy names_cs naming program target strict args =
+fun present_code_for thy target some_width module_name args naming program (names, selects) =
+  present selects (mount_serializer thy target some_width (assert_module_name module_name) args naming program names);
+
+fun check_code_for thy target strict args naming program names_cs =
   let
-    val module_name = "Code_Test";
+    val module_name = "Code";
     val { env_var, make_destination, make_command } =
       (#check o the_fundamental thy) target;
     val env_param = getenv env_var;
     fun ext_check env_param p =
       let 
         val destination = make_destination p;
-        val _ = file destination (serialize thy target (SOME 80)
-          (SOME module_name) args naming program names_cs);
-        val cmd = make_command env_param destination module_name;
+        val _ = export (SOME destination) (mount_serializer thy target (SOME 80)
+          module_name args naming program names_cs);
+        val cmd = make_command env_param module_name;
       in if bash ("cd " ^ File.shell_path p ^ " && " ^ cmd ^ " 2>&1") <> 0
         then error ("Code check failed for " ^ target ^ ": " ^ cmd)
         else ()
@@ -344,24 +394,9 @@
     else Cache_IO.with_tmp_dir "Code_Test" (ext_check env_param)
   end;
 
-fun serialize_custom thy (target_name, seri) naming program names =
-  mount_serializer thy (SOME seri) target_name NONE NONE [] naming program names (String [])
-  |> the;
-
 end; (* local *)
 
 
-(* code presentation *)
-
-fun code_of thy target some_width module_name cs names_stmt =
-  let
-    val (names_cs, (naming, program)) = Code_Thingol.consts_program thy false cs;
-  in
-    string (names_stmt naming)
-      (serialize thy target some_width (SOME module_name) [] naming program names_cs)
-  end;
-
-
 (* code generation *)
 
 fun transitivly_non_empty_funs thy naming program =
@@ -379,22 +414,35 @@
       if member (op =) names3 name then SOME c else NONE) (cs2 ~~ names2);
   in union (op =) cs3 cs1 end;
 
+fun prep_destination "" = NONE
+  | prep_destination "-" = NONE
+  | prep_destination s = SOME (Path.explode s);
+
 fun export_code thy cs seris =
   let
     val (names_cs, (naming, program)) = Code_Thingol.consts_program thy false cs;
-    fun mk_seri_dest dest = if dest = "" orelse dest = "-" then export
-      else file (Path.explode dest);
-    val _ = map (fn (((target, module), dest), args) =>
-      (mk_seri_dest dest (serialize thy target NONE module args naming program names_cs))) seris;
+    val _ = map (fn (((target, module_name), some_path), args) =>
+      export_code_for thy some_path target NONE module_name args naming program names_cs) seris;
   in () end;
 
-fun export_code_cmd raw_cs seris thy = export_code thy (read_const_exprs thy raw_cs) seris;
+fun export_code_cmd raw_cs seris thy = export_code thy (read_const_exprs thy raw_cs)
+  ((map o apfst o apsnd) prep_destination seris);
+
+fun produce_code thy cs target some_width some_module_name args =
+  let
+    val (names_cs, (naming, program)) = Code_Thingol.consts_program thy false cs;
+  in produce_code_for thy target some_width some_module_name args naming program names_cs end;
+
+fun present_code thy cs names_stmt target some_width some_module_name args =
+  let
+    val (names_cs, (naming, program)) = Code_Thingol.consts_program thy false cs;
+  in present_code_for thy target some_width some_module_name args naming program (names_cs, names_stmt naming) end;
 
 fun check_code thy cs seris =
   let
     val (names_cs, (naming, program)) = Code_Thingol.consts_program thy false cs;
-    val _ = map (fn ((target, strict), args) => check thy names_cs naming program
-      target strict args) seris;
+    val _ = map (fn ((target, strict), args) =>
+      check_code_for thy target strict args naming program names_cs) seris;
   in () end;
 
 fun check_code_cmd raw_cs seris thy = check_code thy (read_const_exprs thy raw_cs) seris;
@@ -431,21 +479,21 @@
     val change = case some_raw_syn
      of SOME raw_syn => upd (x, prep_syn thy x raw_syn)
       | NONE => del x;
-  in (map_name_syntax target o mapp) change thy end;
+  in (map_symbol_syntax target o mapp) change thy end;
 
-fun gen_add_syntax_class prep_class =
+fun gen_add_class_syntax prep_class =
   gen_add_syntax (apfst o apfst, Symtab.update, Symtab.delete_safe) prep_class ((K o K) I);
 
-fun gen_add_syntax_inst prep_inst =
+fun gen_add_instance_syntax prep_inst =
   gen_add_syntax (apfst o apsnd, Symreltab.update, Symreltab.delete_safe) prep_inst ((K o K) I);
 
-fun gen_add_syntax_tyco prep_tyco =
+fun gen_add_tyco_syntax prep_tyco =
   gen_add_syntax (apsnd o apfst, Symtab.update, Symtab.delete_safe) prep_tyco
     (fn thy => fn tyco => fn syn => if fst syn <> Sign.arity_number thy tyco
       then error ("Number of arguments mismatch in syntax for type constructor " ^ quote tyco)
       else syn);
 
-fun gen_add_syntax_const prep_const =
+fun gen_add_const_syntax prep_const =
   gen_add_syntax (apsnd o apsnd, Symtab.update, Symtab.delete_safe) prep_const
     (fn thy => fn c => fn syn =>
       if Code_Printer.requires_args syn > Code.args_number thy c
@@ -474,15 +522,17 @@
 val add_include = gen_add_include (K I);
 val add_include_cmd = gen_add_include Code.read_const;
 
-fun add_module_alias target (thyname, modlname) =
-  let
-    val xs = Long_Name.explode modlname;
-    val xs' = map (Name.desymbolize true) xs;
-  in if xs' = xs
-    then map_module_alias target (Symtab.update (thyname, modlname))
-    else error ("Invalid module name: " ^ quote modlname ^ "\n"
-      ^ "perhaps try " ^ quote (Long_Name.implode xs'))
-  end;
+fun add_module_alias target (thyname, "") =
+      map_module_alias target (Symtab.delete thyname)
+  | add_module_alias target (thyname, modlname) =
+      let
+        val xs = Long_Name.explode modlname;
+        val xs' = map (Name.desymbolize true) xs;
+      in if xs' = xs
+        then map_module_alias target (Symtab.update (thyname, modlname))
+        else error ("Invalid module name: " ^ quote modlname ^ "\n"
+          ^ "perhaps try " ^ quote (Long_Name.implode xs'))
+      end;
 
 fun gen_allow_abort prep_const raw_c thy =
   let
@@ -509,18 +559,18 @@
 
 in
 
-val add_syntax_class = gen_add_syntax_class cert_class;
-val add_syntax_inst = gen_add_syntax_inst cert_inst;
-val add_syntax_tyco = gen_add_syntax_tyco cert_tyco;
-val add_syntax_const = gen_add_syntax_const (K I);
+val add_class_syntax = gen_add_class_syntax cert_class;
+val add_instance_syntax = gen_add_instance_syntax cert_inst;
+val add_tyco_syntax = gen_add_tyco_syntax cert_tyco;
+val add_const_syntax = gen_add_const_syntax (K I);
 val allow_abort = gen_allow_abort (K I);
 val add_reserved = add_reserved;
 val add_include = add_include;
 
-val add_syntax_class_cmd = gen_add_syntax_class read_class;
-val add_syntax_inst_cmd = gen_add_syntax_inst read_inst;
-val add_syntax_tyco_cmd = gen_add_syntax_tyco read_tyco;
-val add_syntax_const_cmd = gen_add_syntax_const Code.read_const;
+val add_class_syntax_cmd = gen_add_class_syntax read_class;
+val add_instance_syntax_cmd = gen_add_instance_syntax read_inst;
+val add_tyco_syntax_cmd = gen_add_tyco_syntax read_tyco;
+val add_const_syntax_cmd = gen_add_const_syntax Code.read_const;
 val allow_abort_cmd = gen_allow_abort Code.read_const;
 
 fun parse_args f args =
@@ -541,7 +591,7 @@
       -- ((Parse.$$$ "?" |-- Scan.succeed false) || Scan.succeed true) -- code_expr_argsP))
       >> (fn seris => check_code_cmd raw_cs seris)
     || Scan.repeat (Parse.$$$ inK |-- Parse.name
-       -- Scan.option (Parse.$$$ module_nameK |-- Parse.name)
+       -- Scan.optional (Parse.$$$ module_nameK |-- Parse.name) ""
        -- Scan.optional (Parse.$$$ fileK |-- Parse.name) ""
        -- code_expr_argsP) >> (fn seris => export_code_cmd raw_cs seris)));
 
@@ -550,23 +600,23 @@
 val _ =
   Outer_Syntax.command "code_class" "define code syntax for class" Keyword.thy_decl (
     process_multi_syntax Parse.xname (Scan.option Parse.string)
-    add_syntax_class_cmd);
+    add_class_syntax_cmd);
 
 val _ =
   Outer_Syntax.command "code_instance" "define code syntax for instance" Keyword.thy_decl (
     process_multi_syntax (Parse.xname --| Parse.$$$ "::" -- Parse.xname)
       (Scan.option (Parse.minus >> K ()))
-    add_syntax_inst_cmd);
+    add_instance_syntax_cmd);
 
 val _ =
   Outer_Syntax.command "code_type" "define code syntax for type constructor" Keyword.thy_decl (
     process_multi_syntax Parse.xname Code_Printer.parse_tyco_syntax
-    add_syntax_tyco_cmd);
+    add_tyco_syntax_cmd);
 
 val _ =
   Outer_Syntax.command "code_const" "define code syntax for constant" Keyword.thy_decl (
     process_multi_syntax Parse.term_group Code_Printer.parse_const_syntax
-    add_syntax_const_cmd);
+    add_const_syntax_cmd);
 
 val _ =
   Outer_Syntax.command "code_reserved" "declare words as reserved for target language"
--- a/src/Tools/Code/code_thingol.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/Code/code_thingol.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -271,9 +271,6 @@
        of SOME (tyco, _) => Codegen.thyname_of_type thy tyco
         | NONE => Codegen.thyname_of_const thy c);
   fun purify_base "==>" = "follows"
-    | purify_base "op &" = "and"
-    | purify_base "op |" = "or"
-    | purify_base "op -->" = "implies"
     | purify_base "op =" = "eq"
     | purify_base s = Name.desymbolize false s;
   fun namify thy get_basename get_thyname name =
--- a/src/Tools/Code/lib/Tools/codegen	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/Code/lib/Tools/codegen	Wed Sep 01 17:19:47 2010 +0200
@@ -53,13 +53,13 @@
 
 if [ "$QUICK_AND_DIRTY" -eq 1 ]
 then
-  QND_CMD="set"
+  QND_FLAG="true"
 else
-  QND_CMD="reset"
+  QND_FLAG="false"
 fi
 
 CTXT_CMD="ML_Context.eval_text_in (SOME (ProofContext.init_global (Thy_Info.get_theory \"HOL\"))) false Position.none \"Code_Target.shell_command thyname cmd\";"
 
-FULL_CMD="Unsynchronized.$QND_CMD quick_and_dirty; val thyname = \"$THY\"; val cmd = \"$CODE_CMD\"; $CTXT_CMD"
+FULL_CMD="quick_and_dirty := QND_FLAG; val thyname = \"$THY\"; val cmd = \"$CODE_CMD\"; $CTXT_CMD"
 
 "$ISABELLE_PROCESS" -q -e "$FULL_CMD" "$IMAGE" || exit 1
--- a/src/Tools/Code_Generator.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/Code_Generator.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -17,10 +17,11 @@
   "~~/src/Tools/Code/code_simp.ML"
   "~~/src/Tools/Code/code_printer.ML"
   "~~/src/Tools/Code/code_target.ML"
+  "~~/src/Tools/Code/code_namespace.ML"
   "~~/src/Tools/Code/code_ml.ML"
-  "~~/src/Tools/Code/code_eval.ML"
   "~~/src/Tools/Code/code_haskell.ML"
   "~~/src/Tools/Code/code_scala.ML"
+  "~~/src/Tools/Code/code_eval.ML"
   "~~/src/Tools/nbe.ML"
 begin
 
@@ -28,9 +29,9 @@
   Code_Preproc.setup
   #> Code_Simp.setup
   #> Code_ML.setup
-  #> Code_Eval.setup
   #> Code_Haskell.setup
   #> Code_Scala.setup
+  #> Code_Eval.setup
   #> Nbe.setup
   #> Quickcheck.setup
 *}
--- a/src/Tools/jEdit/dist-template/etc/settings	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/jEdit/dist-template/etc/settings	Wed Sep 01 17:19:47 2010 +0200
@@ -4,8 +4,8 @@
 JEDIT_SETTINGS="$ISABELLE_HOME_USER/jedit"
 
 JEDIT_OPTIONS="-reuseview -noserver -nobackground"
-JEDIT_JAVA_OPTIONS="-Xms128m -Xmx512m -Xss2m -Dactors.corePoolSize=4"
-#JEDIT_JAVA_OPTIONS="-server -Xms128m -Xmx1024m -Xss4m -Dactors.corePoolSize=8"
+JEDIT_JAVA_OPTIONS="-Xms128m -Xmx512m -Xss2m -Dactors.corePoolSize=4 -Dactors.enableForkJoin=false"
+#JEDIT_JAVA_OPTIONS="-server -Xms128m -Xmx1024m -Xss4m -Dactors.corePoolSize=8 -Dactors.enableForkJoin=false"
 JEDIT_SYSTEM_OPTIONS="-Dapple.laf.useScreenMenuBar=true -Dcom.apple.mrj.application.apple.menu.about.name=Isabelle/jEdit"
 
 JEDIT_STYLE_SHEETS="$ISABELLE_HOME/lib/html/isabelle.css:$JEDIT_HOME/etc/isabelle-jedit.css:$ISABELLE_HOME_USER/etc/isabelle.css:$ISABELLE_HOME_USER/etc/isabelle-jedit.css"
--- a/src/Tools/jEdit/plugin/Isabelle.props	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/jEdit/plugin/Isabelle.props	Wed Sep 01 17:19:47 2010 +0200
@@ -29,6 +29,8 @@
 options.isabelle.relative-font-size=100
 options.isabelle.tooltip-font-size.title=Tooltip Font Size
 options.isabelle.tooltip-font-size=10
+options.isabelle.tooltip-dismiss-delay.title=Tooltip Dismiss Delay (global)
+options.isabelle.tooltip-dismiss-delay=8000
 options.isabelle.startup-timeout=10000
 
 #menu actions
--- a/src/Tools/jEdit/src/jedit/document_model.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/jEdit/src/jedit/document_model.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -57,9 +57,18 @@
 
     /* line context */
 
-    private val rule_set = new ParserRuleSet("isabelle", "MAIN")
-    class LineContext(val line: Int, prev: LineContext)
-      extends TokenMarker.LineContext(rule_set, prev)
+    private val dummy_rules = new ParserRuleSet("isabelle", "MAIN")
+
+    class Line_Context(val line: Int, prev: Line_Context)
+      extends TokenMarker.LineContext(dummy_rules, prev)
+    {
+      override def hashCode: Int = line
+      override def equals(that: Any) =
+        that match {
+          case other: Line_Context => line == other.line
+          case _ => false
+        }
+    }
 
 
     /* mapping to jEdit token types */
@@ -181,19 +190,6 @@
 
 class Document_Model(val session: Session, val buffer: Buffer, val thy_name: String)
 {
-  /* visible line end */
-
-  // simplify slightly odd result of TextArea.getLineEndOffset
-  // NB: jEdit already normalizes \r\n and \r to \n
-  def visible_line_end(start: Text.Offset, end: Text.Offset): Text.Offset =
-  {
-    val end1 = end - 1
-    if (start <= end1 && end1 < buffer.getLength &&
-        buffer.getSegment(end1, 1).charAt(0) == '\n') end1
-    else end
-  }
-
-
   /* pending text edits */
 
   object pending_edits  // owned by Swing thread
@@ -256,18 +252,15 @@
     override def markTokens(prev: TokenMarker.LineContext,
         handler: TokenHandler, line_segment: Segment): TokenMarker.LineContext =
     {
-      // FIXME proper synchronization / thread context (!??)
-      val snapshot = Swing_Thread.now { Document_Model.this.snapshot() }
+      Isabelle.swing_buffer_lock(buffer) {
+        val snapshot = Document_Model.this.snapshot()
 
-      Isabelle.buffer_read_lock(buffer) {
-        val previous = prev.asInstanceOf[Document_Model.Token_Markup.LineContext]
+        val previous = prev.asInstanceOf[Document_Model.Token_Markup.Line_Context]
         val line = if (prev == null) 0 else previous.line + 1
-        val context = new Document_Model.Token_Markup.LineContext(line, previous)
+        val context = new Document_Model.Token_Markup.Line_Context(line, previous)
 
         val start = buffer.getLineStartOffset(line)
         val stop = start + line_segment.count
-        val range = Text.Range(start, stop)
-        val former_range = snapshot.revert(range)
 
         /* FIXME
         for (text_area <- Isabelle.jedit_text_areas(buffer)
@@ -290,27 +283,16 @@
             Document_Model.Token_Markup.token_style(name)
         }
 
-        var next_x = start
-        for {
-          (command, command_start) <- snapshot.node.command_range(former_range)
-          info <- snapshot.state(command).markup.
-            select((former_range - command_start).restrict(command.range))(token_markup)(Token.NULL)
-          val Text.Range(abs_start, abs_stop) = snapshot.convert(info.range + command_start)
-          if abs_stop > start && abs_start < stop  // FIXME abs_range overlaps range (redundant!?)
+        var last = start
+        for (token <- snapshot.select_markup(Text.Range(start, stop))(token_markup)(Token.NULL)) {
+          val Text.Range(token_start, token_stop) = token.range
+          if (last < token_start)
+            handle_token(Token.COMMENT1, last - start, token_start - last)
+          handle_token(token.info, token_start - start, token_stop - token_start)
+          last = token_stop
         }
-        {
-          val token_start = (abs_start - start) max 0
-          val token_length =
-            (abs_stop - abs_start) -
-            ((start - abs_start) max 0) -
-            ((abs_stop - stop) max 0)
-          if (start + token_start > next_x)  // FIXME ??
-            handle_token(Token.COMMENT1, next_x - start, start + token_start - next_x)
-          handle_token(info.info, token_start, token_length)
-          next_x = start + token_start + token_length
-        }
-        if (next_x < stop)  // FIXME ??
-          handle_token(Token.COMMENT1, next_x - start, stop - next_x)
+        if (last < stop)
+          handle_token(Token.COMMENT1, last - start, stop - last)
 
         handle_token(Token.END, line_segment.count, 0)
         handler.setLineContext(context)
--- a/src/Tools/jEdit/src/jedit/document_view.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/jEdit/src/jedit/document_view.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -97,35 +97,52 @@
   }
 
 
+  /* visible line ranges */
+
+  // simplify slightly odd result of TextArea.getScreenLineEndOffset etc.
+  // NB: jEdit already normalizes \r\n and \r to \n
+  def proper_line_range(start: Text.Offset, end: Text.Offset): Text.Range =
+  {
+    val stop = if (start < end) end - 1 else end min model.buffer.getLength
+    Text.Range(start, stop)
+  }
+
+  def screen_lines_range(): Text.Range =
+  {
+    val start = text_area.getScreenLineStartOffset(0)
+    val raw_end = text_area.getScreenLineEndOffset(text_area.getVisibleLines - 1 max 0)
+    proper_line_range(start, if (raw_end >= 0) raw_end else model.buffer.getLength)
+  }
+
+
   /* commands_changed_actor */
 
   private val commands_changed_actor = actor {
     loop {
       react {
         case Session.Commands_Changed(changed) =>
-          Swing_Thread.now {
-            // FIXME cover doc states as well!!?
+          val buffer = model.buffer
+          Isabelle.swing_buffer_lock(buffer) {
             val snapshot = model.snapshot()
-            val buffer = model.buffer
-            Isabelle.buffer_read_lock(buffer) {
-              if (changed.exists(snapshot.node.commands.contains)) {
-                var visible_change = false
+
+            if (changed.exists(snapshot.node.commands.contains))
+              overview.repaint()
 
-                for ((command, start) <- snapshot.node.command_starts) {
-                  if (changed(command)) {
-                    val stop = start + command.length
-                    val line1 = buffer.getLineOfOffset(snapshot.convert(start))
-                    val line2 = buffer.getLineOfOffset(snapshot.convert(stop))
-                    if (line2 >= text_area.getFirstLine &&
-                        line1 <= text_area.getFirstLine + text_area.getVisibleLines)
-                      visible_change = true
-                    text_area.invalidateLineRange(line1, line2)
-                  }
-                }
-                if (visible_change) model.buffer.propertiesChanged()
+            val visible_range = screen_lines_range()
+            val visible_cmds = snapshot.node.command_range(snapshot.revert(visible_range)).map(_._1)
+            if (visible_cmds.exists(changed)) {
+              for {
+                line <- 0 until text_area.getVisibleLines
+                val start = text_area.getScreenLineStartOffset(line) if start >= 0
+                val end = text_area.getScreenLineEndOffset(line) if end >= 0
+                val range = proper_line_range(start, end)
+                val line_cmds = snapshot.node.command_range(snapshot.revert(range)).map(_._1)
+                if line_cmds.exists(changed)
+              } text_area.invalidateScreenLineRange(line, line)
 
-                overview.repaint()  // FIXME paint here!?
-              }
+              // FIXME danger of deadlock!?
+              // FIXME potentially slow!?
+              model.buffer.propertiesChanged()
             }
           }
 
@@ -141,71 +158,43 @@
   {
     override def paintScreenLineRange(gfx: Graphics2D,
       first_line: Int, last_line: Int, physical_lines: Array[Int],
-      start: Array[Int], end: Array[Int], y0: Int, line_height: Int)
+      start: Array[Int], end: Array[Int], y: Int, line_height: Int)
     {
-      Swing_Thread.assert()
-
-      val snapshot = model.snapshot()
-
-      val command_range: Iterable[(Command, Text.Offset)] =
-      {
-        val range = snapshot.node.command_range(snapshot.revert(start(0)))
-        if (range.hasNext) {
-          val (cmd0, start0) = range.next
-          new Iterable[(Command, Text.Offset)] {
-            def iterator =
-              Document.Node.command_starts(snapshot.node.commands.iterator(cmd0), start0)
+      Isabelle.swing_buffer_lock(model.buffer) {
+        val snapshot = model.snapshot()
+        val saved_color = gfx.getColor
+        try {
+          for (i <- 0 until physical_lines.length) {
+            if (physical_lines(i) != -1) {
+              val line_range = proper_line_range(start(i), end(i))
+              val cmds = snapshot.node.command_range(snapshot.revert(line_range))
+              for ((command, command_start) <- cmds if !command.is_ignored) {
+                val range = line_range.restrict(snapshot.convert(command.range + command_start))
+                val p = text_area.offsetToXY(range.start)
+                val q = text_area.offsetToXY(range.stop)
+                if (p != null && q != null) {
+                  gfx.setColor(Document_View.choose_color(snapshot, command))
+                  gfx.fillRect(p.x, y + i * line_height, q.x - p.x, line_height)
+                }
+              }
+            }
           }
         }
-        else Iterable.empty
+        finally { gfx.setColor(saved_color) }
       }
-
-      val saved_color = gfx.getColor
-      try {
-        var y = y0
-        for (i <- 0 until physical_lines.length) {
-          if (physical_lines(i) != -1) {
-            val line_start = start(i)
-            val line_end = model.visible_line_end(line_start, end(i))
-
-            val a = snapshot.revert(line_start)
-            val b = snapshot.revert(line_end)
-            val cmds = command_range.iterator.
-              dropWhile { case (cmd, c) => c + cmd.length <= a } .
-              takeWhile { case (_, c) => c < b }
-
-            for ((command, command_start) <- cmds if !command.is_ignored) {
-              val p =
-                text_area.offsetToXY(line_start max snapshot.convert(command_start))
-              val q =
-                text_area.offsetToXY(line_end min snapshot.convert(command_start + command.length))
-              assert(p.y == q.y)
-              gfx.setColor(Document_View.choose_color(snapshot, command))
-              gfx.fillRect(p.x, y, q.x - p.x, line_height)
-            }
-          }
-          y += line_height
-        }
-      }
-      finally { gfx.setColor(saved_color) }
     }
 
     override def getToolTipText(x: Int, y: Int): String =
     {
-      Swing_Thread.assert()
-
-      val snapshot = model.snapshot()
-      val offset = snapshot.revert(text_area.xyToOffset(x, y))
-      snapshot.node.command_at(offset) match {
-        case Some((command, command_start)) =>
-          // FIXME Isar_Document.Tooltip extractor
-          (snapshot.state(command).markup.select(Text.Range(offset, offset + 1) - command_start) {
+      Isabelle.swing_buffer_lock(model.buffer) {
+        val snapshot = model.snapshot()
+        val offset = text_area.xyToOffset(x, y)
+        val markup =
+          snapshot.select_markup(Text.Range(offset, offset + 1)) {
             case Text.Info(range, XML.Elem(Markup(Markup.ML_TYPING, _), body)) =>
-              val typing =
-                Pretty.block(XML.Text(command.source(range) + " : ") :: Pretty.Break(1) :: body)
-              Isabelle.tooltip(Pretty.string_of(List(typing), margin = 40))
-          } { null }).head.info
-        case None => null
+              Isabelle.tooltip(Pretty.string_of(List(Pretty.block(body)), margin = 40))
+          } { null }
+        if (markup.hasNext) markup.next.info else null
       }
     }
   }
@@ -268,7 +257,7 @@
       super.paintComponent(gfx)
       Swing_Thread.assert()
       val buffer = model.buffer
-      Isabelle.buffer_read_lock(buffer) {
+      Isabelle.buffer_lock(buffer) {
         val snapshot = model.snapshot()
         val saved_color = gfx.getColor  // FIXME needed!?
         try {
--- a/src/Tools/jEdit/src/jedit/isabelle_hyperlinks.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/jEdit/src/jedit/isabelle_hyperlinks.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -29,7 +29,8 @@
 {
   override def click(view: View) = {
     Isabelle.system.source_file(ref_file) match {
-      case None => System.err.println("Could not find source file " + ref_file)  // FIXME ??
+      case None =>
+        Library.error_dialog(view, "File not found", "Could not find source file " + ref_file)
       case Some(file) =>
         jEdit.openFiles(view, file.getParent, Array(file.getName, "+line:" + ref_line))
     }
@@ -40,27 +41,24 @@
 {
   def getHyperlink(buffer: Buffer, buffer_offset: Int): Hyperlink =
   {
-    // FIXME lock buffer (!??)
     Swing_Thread.assert()
-    Document_Model(buffer) match {
-      case Some(model) =>
-        val snapshot = model.snapshot()
-        val offset = snapshot.revert(buffer_offset)
-        snapshot.node.command_at(offset) match {
-          case Some((command, command_start)) =>
-            // FIXME Isar_Document.Hyperlink extractor
-            (snapshot.state(command).markup.select(Text.Range(offset, offset + 1) - command_start) {
+    Isabelle.buffer_lock(buffer) {
+      Document_Model(buffer) match {
+        case Some(model) =>
+          val snapshot = model.snapshot()
+          val markup =
+            snapshot.select_markup(Text.Range(buffer_offset, buffer_offset + 1)) {
+              // FIXME Isar_Document.Hyperlink extractor
               case Text.Info(info_range, XML.Elem(Markup(Markup.ML_REF, _),
                   List(XML.Elem(Markup(Markup.ML_DEF, props), _)))) =>
-                val Text.Range(begin, end) = snapshot.convert(info_range + command_start)
+                val Text.Range(begin, end) = info_range
                 val line = buffer.getLineOfOffset(begin)
-
-                (Position.get_file(props), Position.get_line(props)) match {
+                (Position.File.unapply(props), Position.Line.unapply(props)) match {
                   case (Some(ref_file), Some(ref_line)) =>
                     new External_Hyperlink(begin, end, line, ref_file, ref_line)
                   case _ =>
-                    (Position.get_id(props), Position.get_offset(props)) match {
-                      case (Some(ref_id), Some(ref_offset)) =>
+                    (props, props) match {
+                      case (Position.Id(ref_id), Position.Offset(ref_offset)) =>
                         snapshot.lookup_command(ref_id) match {
                           case Some(ref_cmd) =>
                             snapshot.node.command_start(ref_cmd) match {
@@ -74,10 +72,11 @@
                       case _ => null
                     }
                 }
-            } { null }).head.info
-          case None => null
-        }
-      case None => null
+            } { null }
+          if (markup.hasNext) markup.next.info else null
+
+        case None => null
+      }
     }
   }
 }
--- a/src/Tools/jEdit/src/jedit/isabelle_options.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/jEdit/src/jedit/isabelle_options.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -17,6 +17,7 @@
   private val logic_name = new JComboBox()
   private val relative_font_size = new JSpinner()
   private val tooltip_font_size = new JSpinner()
+  private val tooltip_dismiss_delay = new JSpinner()
 
   private class List_Item(val name: String, val descr: String) {
     def this(name: String) = this(name, name)
@@ -46,6 +47,11 @@
       tooltip_font_size.setValue(Isabelle.Int_Property("tooltip-font-size", 10))
       tooltip_font_size
     })
+
+    addComponent(Isabelle.Property("tooltip-dismiss-delay.title"), {
+      tooltip_dismiss_delay.setValue(Isabelle.Int_Property("tooltip-dismiss-delay", 8000))
+      tooltip_dismiss_delay
+    })
   }
 
   override def _save()
@@ -58,5 +64,8 @@
 
     Isabelle.Int_Property("tooltip-font-size") =
       tooltip_font_size.getValue().asInstanceOf[Int]
+
+    Isabelle.Int_Property("tooltip-dismiss-delay") =
+      tooltip_dismiss_delay.getValue().asInstanceOf[Int]
   }
 }
--- a/src/Tools/jEdit/src/jedit/output_dockable.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/jEdit/src/jedit/output_dockable.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -67,12 +67,12 @@
             case Some(cmd) if !restriction.isDefined || restriction.get.contains(cmd) =>
               val snapshot = doc_view.model.snapshot()
               val filtered_results =
-                snapshot.state(cmd).results filter {
+                snapshot.state(cmd).results.iterator.map(_._2) filter {
                   case XML.Elem(Markup(Markup.TRACING, _), _) => show_tracing
                   case XML.Elem(Markup(Markup.DEBUG, _), _) => show_debug
                   case _ => true
                 }
-              html_panel.render(filtered_results)
+              html_panel.render(filtered_results.toList)
             case _ =>
           }
         case None =>
--- a/src/Tools/jEdit/src/jedit/plugin.scala	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/jEdit/src/jedit/plugin.scala	Wed Sep 01 17:19:47 2010 +0200
@@ -81,6 +81,17 @@
         Int_Property("tooltip-font-size", 10).toString + "px; \">" +  // FIXME proper scaling (!?)
       HTML.encode(text) + "</pre></html>"
 
+  def tooltip_dismiss_delay(): Int =
+    Int_Property("tooltip-dismiss-delay", 8000) max 500
+
+  def setup_tooltips()
+  {
+    Swing_Thread.now {
+      val manager = javax.swing.ToolTipManager.sharedInstance
+      manager.setDismissDelay(tooltip_dismiss_delay())
+    }
+  }
+
 
   /* settings */
 
@@ -118,12 +129,15 @@
   def jedit_text_areas(buffer: JEditBuffer): Iterator[JEditTextArea] =
     jedit_text_areas().filter(_.getBuffer == buffer)
 
-  def buffer_read_lock[A](buffer: JEditBuffer)(body: => A): A =
+  def buffer_lock[A](buffer: JEditBuffer)(body: => A): A =
   {
     try { buffer.readLock(); body }
     finally { buffer.readUnlock() }
   }
 
+  def swing_buffer_lock[A](buffer: JEditBuffer)(body: => A): A =
+    Swing_Thread.now { buffer_lock(buffer) { body } }
+
 
   /* dockable windows */
 
@@ -240,6 +254,8 @@
         Swing_Thread.now {
           for (text_area <- Isabelle.jedit_text_areas if Document_View(text_area).isDefined)
             Document_View(text_area).get.extend_styles()
+
+          Isabelle.setup_tooltips()
         }
 
         Isabelle.session.global_settings.event(Session.Global_Settings)
@@ -253,6 +269,8 @@
     Isabelle.system = new Isabelle_System
     Isabelle.system.install_fonts()
     Isabelle.session = new Session(Isabelle.system)  // FIXME dialog!?
+
+    Isabelle.setup_tooltips()
   }
 
   override def stop()
--- a/src/Tools/nbe.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/nbe.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -388,6 +388,7 @@
       in
         s
         |> traced (fn s => "\n--- code to be evaluated:\n" ^ s)
+        |> pair ""
         |> ML_Context.evaluate ctxt (!trace) univs_cookie
         |> (fn f => f deps_vals)
         |> (fn univs => cs ~~ univs)
--- a/src/Tools/quickcheck.ML	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/Tools/quickcheck.ML	Wed Sep 01 17:19:47 2010 +0200
@@ -78,27 +78,32 @@
 
 datatype test_params = Test_Params of
   { size: int, iterations: int, default_type: typ list, no_assms: bool,
-  expect : expectation, report: bool, quiet : bool};
+    expect : expectation, report: bool, quiet : bool};
 
 fun dest_test_params (Test_Params { size, iterations, default_type, no_assms, expect, report, quiet }) =
   ((size, iterations), ((default_type, no_assms), ((expect, report), quiet)));
+
 fun make_test_params ((size, iterations), ((default_type, no_assms), ((expect, report), quiet))) =
   Test_Params { size = size, iterations = iterations, default_type = default_type,
-                no_assms = no_assms, expect = expect, report = report, quiet = quiet };
+    no_assms = no_assms, expect = expect, report = report, quiet = quiet };
+
 fun map_test_params f (Test_Params { size, iterations, default_type, no_assms, expect, report, quiet }) =
   make_test_params (f ((size, iterations), ((default_type, no_assms), ((expect, report), quiet))));
-fun merge_test_params (Test_Params { size = size1, iterations = iterations1, default_type = default_type1,
-                                     no_assms = no_assms1, expect = expect1, report = report1, quiet = quiet1 },
+
+fun merge_test_params
+ (Test_Params { size = size1, iterations = iterations1, default_type = default_type1,
+    no_assms = no_assms1, expect = expect1, report = report1, quiet = quiet1 },
   Test_Params { size = size2, iterations = iterations2, default_type = default_type2,
-                no_assms = no_assms2, expect = expect2, report = report2, quiet = quiet2 }) =
+    no_assms = no_assms2, expect = expect2, report = report2, quiet = quiet2 }) =
   make_test_params ((Int.max (size1, size2), Int.max (iterations1, iterations2)),
     ((merge (op =) (default_type1, default_type2), no_assms1 orelse no_assms2),
     ((merge_expectation (expect1, expect2), report1 orelse report2), quiet1 orelse quiet2)));
 
 structure Data = Theory_Data
 (
-  type T = (string * (Proof.context -> bool -> term -> int -> term list option * (bool list * bool))) list
-    * test_params;
+  type T =
+    (string * (Proof.context -> bool -> term -> int -> term list option * (bool list * bool))) list
+      * test_params;
   val empty = ([], Test_Params
     { size = 10, iterations = 100, default_type = [], no_assms = false,
       expect = No_Expectation, report = false, quiet = false});
--- a/src/ZF/ZF.thy	Thu Aug 26 13:15:37 2010 +0200
+++ b/src/ZF/ZF.thy	Wed Sep 01 17:19:47 2010 +0200
@@ -10,7 +10,7 @@
 uses "~~/src/Tools/misc_legacy.ML"
 begin
 
-ML {* Unsynchronized.reset eta_contract *}
+ML {* eta_contract := false *}
 
 typedecl i
 arities  i :: "term"