merged
authorwenzelm
Thu, 19 Aug 2010 17:41:52 +0200
changeset 38559 befdd6833ec0
parent 38558 32ad17fe2b9c (diff)
parent 38486 f5bbfc019937 (current diff)
child 38560 004c35739d75
merged
NEWS
lib/scripts/run-polyml-5.0
src/Pure/ML-Systems/compiler_polyml-5.0.ML
src/Pure/ML-Systems/polyml-5.0.ML
src/Pure/ML-Systems/polyml-5.1.ML
src/Pure/PIDE/markup_node.scala
src/Pure/System/isar_document.ML
src/Pure/System/isar_document.scala
--- a/NEWS	Thu Aug 19 17:40:44 2010 +0200
+++ b/NEWS	Thu Aug 19 17:41:52 2010 +0200
@@ -35,6 +35,18 @@
 
 *** HOL ***
 
+* Records: logical foundation type for records do not carry a '_type' suffix
+any longer.  INCOMPATIBILITY.
+
+* Code generation for records: more idiomatic representation of record types.
+Warning: records are not covered by ancient SML code generation any longer.
+INCOMPATIBILITY.  In cases of need, a suitable rep_datatype declaration
+helps to succeed then:
+
+  record 'a foo = ...
+  ...
+  rep_datatype foo_ext ...
+
 * Session Imperative_HOL: revamped, corrected dozens of inadequacies.
 INCOMPATIBILITY.
 
@@ -74,9 +86,20 @@
 * Some previously unqualified names have been qualified:
 
   types
+    bool ~> HOL.bool
     nat ~> Nat.nat
 
   constants
+    Trueprop ~> HOL.Trueprop
+    True ~> HOL.True
+    False ~> HOL.False
+    Not ~> HOL.Not
+    The ~> HOL.The
+    All ~> HOL.All
+    Ex ~> HOL.Ex
+    Ex1 ~> HOL.Ex1
+    Let ~> HOL.Let
+    If ~> HOL.If
     Ball ~> Set.Ball
     Bex ~> Set.Bex
     Suc ~> Nat.Suc
@@ -90,7 +113,7 @@
 INCOMPATIBILITY.
 
 * Removed simplifier congruence rule of "prod_case", as has for long
-been the case with "split".
+been the case with "split".  INCOMPATIBILITY.
 
 * Datatype package: theorems generated for executable equality
 (class eq) carry proper names and are treated as default code
@@ -99,8 +122,8 @@
 * List.thy: use various operations from the Haskell prelude when
 generating Haskell code.
 
-* code_simp.ML: simplification with rules determined by
-code generator.
+* code_simp.ML and method code_simp: simplification with rules determined
+by code generator.
 
 * code generator: do not print function definitions for case
 combinators any longer.
@@ -131,6 +154,16 @@
 similar to inductive_cases.
 
 
+*** FOL ***
+
+* All constant names are now qualified.  INCOMPATIBILITY.
+
+
+*** ZF ***
+
+* All constant names are now qualified.  INCOMPATIBILITY.
+
+
 *** ML ***
 
 * ML antiquotations @{theory} and @{theory_ref} refer to named
--- a/doc-src/Codegen/IsaMakefile	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/IsaMakefile	Thu Aug 19 17:41:52 2010 +0200
@@ -24,7 +24,7 @@
 Thy: $(THY)
 
 $(THY): Thy/ROOT.ML Thy/*.thy ../antiquote_setup.ML ../more_antiquote.ML
-	@$(USEDIR) -m no_brackets -m iff HOL Thy
+	@$(USEDIR) -m no_brackets -m iff HOL-Library Thy
 	@rm -f Thy/document/isabelle.sty Thy/document/isabellesym.sty \
 	 Thy/document/pdfsetup.sty Thy/document/session.tex
 
--- a/doc-src/Codegen/Thy/Adaptation.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/Thy/Adaptation.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -194,12 +194,12 @@
   (SML "true" and "false" and "_ andalso _")
 
 text {*
-  \noindent The @{command code_type} command takes a type constructor
+  \noindent The @{command_def code_type} command takes a type constructor
   as arguments together with a list of custom serialisations.  Each
   custom serialisation starts with a target language identifier
   followed by an expression, which during code serialisation is
   inserted whenever the type constructor would occur.  For constants,
-  @{command code_const} implements the corresponding mechanism.  Each
+  @{command_def code_const} implements the corresponding mechanism.  Each
   ``@{verbatim "_"}'' in a serialisation expression is treated as a
   placeholder for the type constructor's (the constant's) arguments.
 *}
@@ -226,7 +226,7 @@
   avoided to be used for new declarations.  Initially, this table
   typically contains the keywords of the target language.  It can be
   extended manually, thus avoiding accidental overwrites, using the
-  @{command "code_reserved"} command:
+  @{command_def "code_reserved"} command:
 *}
 
 code_reserved %quote "\<SML>" bool true false andalso
@@ -274,7 +274,7 @@
   For convenience, the default @{text HOL} setup for @{text Haskell}
   maps the @{class eq} class to its counterpart in @{text Haskell},
   giving custom serialisations for the class @{class eq} (by command
-  @{command code_class}) and its operation @{const HOL.eq}
+  @{command_def code_class}) and its operation @{const HOL.eq}
 *}
 
 code_class %quotett eq
@@ -307,7 +307,7 @@
 text {*
   \noindent The code generator would produce an additional instance,
   which of course is rejected by the @{text Haskell} compiler.  To
-  suppress this additional instance, use @{text "code_instance"}:
+  suppress this additional instance, use @{command_def "code_instance"}:
 *}
 
 code_instance %quotett bar :: eq
@@ -318,7 +318,7 @@
 
 text {*
   In rare cases it is necessary to \emph{enrich} the context of a
-  target language; this is accomplished using the @{command
+  target language; this is accomplished using the @{command_def
   "code_include"} command:
 *}
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/Thy/Evaluation.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -0,0 +1,104 @@
+theory Evaluation
+imports Setup
+begin
+
+section {* Evaluation *}
+
+text {* Introduction *}
+
+
+subsection {* Evaluation techniques *}
+
+text {* simplifier *}
+
+text {* nbe *}
+
+text {* eval target: SML standalone vs. Isabelle/SML, example, soundness *}
+
+
+subsection {* Dynamic evaluation *}
+
+text {* value (three variants) *}
+
+text {* methods (three variants) *}
+
+text {* corresponding ML interfaces *}
+
+
+subsection {* Static evaluation *}
+
+text {* code_simp, nbe (tbd), Eval (tbd, in simple fashion) *}
+
+text {* hand-written: code antiquotation *}
+
+
+subsection {* Hybrid techniques *}
+
+text {* code reflect runtime *}
+
+text {* code reflect external *}
+
+
+text {* FIXME here the old sections follow *}
+
+subsection {* Evaluation oracle *}
+
+text {*
+  Code generation may also be used to \emph{evaluate} expressions
+  (using @{text SML} as target language of course).
+  For instance, the @{command_def value} reduces an expression to a
+  normal form with respect to the underlying code equations:
+*}
+
+value %quote "42 / (12 :: rat)"
+
+text {*
+  \noindent will display @{term "7 / (2 :: rat)"}.
+
+  The @{method eval} method tries to reduce a goal by code generation to @{term True}
+  and solves it in that case, but fails otherwise:
+*}
+
+lemma %quote "42 / (12 :: rat) = 7 / 2"
+  by %quote eval
+
+text {*
+  \noindent The soundness of the @{method eval} method depends crucially 
+  on the correctness of the code generator;  this is one of the reasons
+  why you should not use adaptation (see \secref{sec:adaptation}) frivolously.
+*}
+
+
+subsubsection {* Code antiquotation *}
+
+text {*
+  In scenarios involving techniques like reflection it is quite common
+  that code generated from a theory forms the basis for implementing
+  a proof procedure in @{text SML}.  To facilitate interfacing of generated code
+  with system code, the code generator provides a @{text code} antiquotation:
+*}
+
+datatype %quote form = T | F | And form form | Or form form (*<*)
+
+(*>*) ML %quotett {*
+  fun eval_form @{code T} = true
+    | eval_form @{code F} = false
+    | eval_form (@{code And} (p, q)) =
+        eval_form p andalso eval_form q
+    | eval_form (@{code Or} (p, q)) =
+        eval_form p orelse eval_form q;
+*}
+
+text {*
+  \noindent @{text code} takes as argument the name of a constant;  after the
+  whole @{text SML} is read, the necessary code is generated transparently
+  and the corresponding constant names are inserted.  This technique also
+  allows to use pattern matching on constructors stemming from compiled
+  @{text "datatypes"}.
+
+  For a less simplistic example, theory @{theory Ferrack} is
+  a good reference.
+*}
+
+
+end
--- a/doc-src/Codegen/Thy/Foundations.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/Thy/Foundations.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -121,9 +121,10 @@
   \secref{eff_nat}) uses this interface.
 
   \noindent The current setup of the preprocessor may be inspected
-  using the @{command print_codeproc} command.  @{command code_thms}
-  (see \secref{sec:equations}) provides a convenient mechanism to
-  inspect the impact of a preprocessor setup on code equations.
+  using the @{command_def print_codeproc} command.  @{command_def
+  code_thms} (see \secref{sec:equations}) provides a convenient
+  mechanism to inspect the impact of a preprocessor setup on code
+  equations.
 
   \begin{warn}
     Attribute @{attribute code_unfold} also applies to the
@@ -175,9 +176,8 @@
   equations (before preprocessing) and code equations (after
   preprocessing).
 
-  The first can be listed (among other data) using the @{command
-  print_codesetup} command (later on in \secref{sec:refinement}
-  it will be shown how these code equations can be changed).
+  The first can be listed (among other data) using the @{command_def
+  print_codesetup} command.
 
   The code equations after preprocessing are already are blueprint of
   the generated program and can be inspected using the @{command
@@ -190,7 +190,7 @@
   \noindent This prints a table with the code equations for @{const
   dequeue}, including \emph{all} code equations those equations depend
   on recursively.  These dependencies themselves can be visualized using
-  the @{command code_deps} command.
+  the @{command_def code_deps} command.
 *}
 
 
@@ -280,7 +280,7 @@
   of as function definitions which always fail,
   since there is never a successful pattern match on the left hand
   side.  In order to categorise a constant into that category
-  explicitly, use @{command "code_abort"}:
+  explicitly, use @{command_def "code_abort"}:
 *}
 
 code_abort %quote empty_queue
--- a/doc-src/Codegen/Thy/Further.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/Thy/Further.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -75,7 +75,7 @@
 end %quote
 
 text {*
-  After an interpretation of this locale (say, @{command
+  After an interpretation of this locale (say, @{command_def
   interpretation} @{text "fun_power:"} @{term [source] "power (\<lambda>n (f ::
   'a \<Rightarrow> 'a). f ^^ n)"}), one would expect to have a constant @{text
   "fun_power.powers :: nat list \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"} for which code
@@ -126,91 +126,27 @@
 *}
 
 
-subsection {* Evaluation oracle *}
-
-text {*
-  Code generation may also be used to \emph{evaluate} expressions
-  (using @{text SML} as target language of course).
-  For instance, the @{command value} reduces an expression to a
-  normal form with respect to the underlying code equations:
-*}
-
-value %quote "42 / (12 :: rat)"
-
-text {*
-  \noindent will display @{term "7 / (2 :: rat)"}.
-
-  The @{method eval} method tries to reduce a goal by code generation to @{term True}
-  and solves it in that case, but fails otherwise:
-*}
-
-lemma %quote "42 / (12 :: rat) = 7 / 2"
-  by %quote eval
-
-text {*
-  \noindent The soundness of the @{method eval} method depends crucially 
-  on the correctness of the code generator;  this is one of the reasons
-  why you should not use adaptation (see \secref{sec:adaptation}) frivolously.
-*}
-
-
-subsection {* Building evaluators *}
-
-text {*
-  FIXME
-*}
-
-subsubsection {* Code antiquotation *}
-
-text {*
-  In scenarios involving techniques like reflection it is quite common
-  that code generated from a theory forms the basis for implementing
-  a proof procedure in @{text SML}.  To facilitate interfacing of generated code
-  with system code, the code generator provides a @{text code} antiquotation:
-*}
-
-datatype %quote form = T | F | And form form | Or form form (*<*)
-
-(*>*) ML %quotett {*
-  fun eval_form @{code T} = true
-    | eval_form @{code F} = false
-    | eval_form (@{code And} (p, q)) =
-        eval_form p andalso eval_form q
-    | eval_form (@{code Or} (p, q)) =
-        eval_form p orelse eval_form q;
-*}
-
-text {*
-  \noindent @{text code} takes as argument the name of a constant;  after the
-  whole @{text SML} is read, the necessary code is generated transparently
-  and the corresponding constant names are inserted.  This technique also
-  allows to use pattern matching on constructors stemming from compiled
-  @{text "datatypes"}.
-
-  For a less simplistic example, theory @{theory Ferrack} is
-  a good reference.
-*}
-
-
 subsection {* ML system interfaces \label{sec:ml} *}
 
 text {*
-  Since the code generator framework not only aims to provide
-  a nice Isar interface but also to form a base for
-  code-generation-based applications, here a short
-  description of the most important ML interfaces.
+  Since the code generator framework not only aims to provide a nice
+  Isar interface but also to form a base for code-generation-based
+  applications, here a short description of the most fundamental ML
+  interfaces.
 *}
 
 subsubsection {* Managing executable content *}
 
 text %mlref {*
   \begin{mldecls}
+  @{index_ML Code.read_const: "theory -> string -> string"} \\
   @{index_ML Code.add_eqn: "thm -> theory -> theory"} \\
   @{index_ML Code.del_eqn: "thm -> theory -> theory"} \\
   @{index_ML Code_Preproc.map_pre: "(simpset -> simpset) -> theory -> theory"} \\
   @{index_ML Code_Preproc.map_post: "(simpset -> simpset) -> theory -> theory"} \\
-  @{index_ML Code_Preproc.add_functrans: "string * (theory -> (thm * bool) list -> (thm * bool) list option)
-    -> theory -> theory"} \\
+  @{index_ML Code_Preproc.add_functrans: "
+    string * (theory -> (thm * bool) list -> (thm * bool) list option)
+      -> theory -> theory"} \\
   @{index_ML Code_Preproc.del_functrans: "string -> theory -> theory"} \\
   @{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\
   @{index_ML Code.get_type: "theory -> string
@@ -220,6 +156,9 @@
 
   \begin{description}
 
+  \item @{ML Code.read_const}~@{text thy}~@{text s}
+     reads a constant as a concrete term expression @{text s}.
+
   \item @{ML Code.add_eqn}~@{text "thm"}~@{text "thy"} adds function
      theorem @{text "thm"} to executable content.
 
@@ -252,21 +191,6 @@
   \end{description}
 *}
 
-subsubsection {* Auxiliary *}
-
-text %mlref {*
-  \begin{mldecls}
-  @{index_ML Code.read_const: "theory -> string -> string"}
-  \end{mldecls}
-
-  \begin{description}
-
-  \item @{ML Code.read_const}~@{text thy}~@{text s}
-     reads a constant as a concrete term expression @{text s}.
-
-  \end{description}
-
-*}
 
 subsubsection {* Data depending on the theory's executable content *}
 
--- a/doc-src/Codegen/Thy/Inductive_Predicate.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/Thy/Inductive_Predicate.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -16,7 +16,7 @@
 section {* Inductive Predicates \label{sec:inductive} *}
 
 text {*
-  The @{text predicate_compiler} is an extension of the code generator
+  The @{text "predicate compiler"} is an extension of the code generator
   which turns inductive specifications into equational ones, from
   which in turn executable code can be generated.  The mechanisms of
   this compiler are described in detail in
--- a/doc-src/Codegen/Thy/Introduction.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/Thy/Introduction.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -34,11 +34,12 @@
 subsection {* A quick start with the Isabelle/HOL toolbox \label{sec:queue_example} *}
 
 text {*
-  In a HOL theory, the @{command datatype} and @{command
-  definition}/@{command primrec}/@{command fun} declarations form the
-  core of a functional programming language.  By default equational
-  theorems stemming from those are used for generated code, therefore
-  \qt{naive} code generation can proceed without further ado.
+  In a HOL theory, the @{command_def datatype} and @{command_def
+  definition}/@{command_def primrec}/@{command_def fun} declarations
+  form the core of a functional programming language.  By default
+  equational theorems stemming from those are used for generated code,
+  therefore \qt{naive} code generation can proceed without further
+  ado.
 
   For example, here a simple \qt{implementation} of amortised queues:
 *}
@@ -71,12 +72,12 @@
 text %quote {*@{code_stmts empty enqueue dequeue (SML)}*}
 
 text {*
-  \noindent The @{command export_code} command takes a space-separated
-  list of constants for which code shall be generated; anything else
-  needed for those is added implicitly.  Then follows a 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
+  \noindent The @{command_def export_code} command takes a
+  space-separated list of constants for which code shall be generated;
+  anything else needed for those is added implicitly.  Then follows a
+  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
   module name (with extension @{text ".hs"}) is written:
--- a/doc-src/Codegen/Thy/ROOT.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/Thy/ROOT.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -5,5 +5,6 @@
 use_thy "Foundations";
 use_thy "Refinement";
 use_thy "Inductive_Predicate";
+use_thy "Evaluation";
 use_thy "Adaptation";
 use_thy "Further";
--- a/doc-src/Codegen/Thy/Refinement.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/Thy/Refinement.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -116,7 +116,7 @@
   "{\<alpha>\<^isub>1, \<dots>, \<alpha>\<^isub>n}"} is exactly the set of \emph{all} type variables in
   @{text "\<tau>"}; then @{text "\<kappa>"} is its corresponding datatype.  The
   HOL datatype package by default registers any new datatype with its
-  constructors, but this may be changed using @{command
+  constructors, but this may be changed using @{command_def
   code_datatype}; the currently chosen constructors can be inspected
   using the @{command print_codesetup} command.
 
@@ -158,7 +158,7 @@
 text {*
   The same techniques can also be applied to types which are not
   specified as datatypes, e.g.~type @{typ int} is originally specified
-  as quotient type by means of @{command typedef}, but for code
+  as quotient type by means of @{command_def typedef}, but for code
   generation constants allowing construction of binary numeral values
   are used as constructors for @{typ int}.
 
@@ -170,7 +170,95 @@
 subsection {* Datatype refinement involving invariants *}
 
 text {*
-  FIXME
+  Datatype representation involving invariants require a dedicated
+  setup for the type and its primitive operations.  As a running
+  example, we implement a type @{text "'a dlist"} of list consisting
+  of distinct elements.
+
+  The first step is to decide on which representation the abstract
+  type (in our example @{text "'a dlist"}) should be implemented.
+  Here we choose @{text "'a list"}.  Then a conversion from the concrete
+  type to the abstract type must be specified, here:
+*}
+
+text %quote {*
+  @{term_type Dlist}
+*}
+
+text {*
+  \noindent Next follows the specification of a suitable \emph{projection},
+  i.e.~a conversion from abstract to concrete type:
+*}
+
+text %quote {*
+  @{term_type list_of_dlist}
+*}
+
+text {*
+  \noindent This projection must be specified such that the following
+  \emph{abstract datatype certificate} can be proven:
+*}
+
+lemma %quote [code abstype]:
+  "Dlist (list_of_dlist dxs) = dxs"
+  by (fact Dlist_list_of_dlist)
+
+text {*
+  \noindent Note that so far the invariant on representations
+  (@{term_type distinct}) has never been mentioned explicitly:
+  the invariant is only referred to implicitly: all values in
+  set @{term "{xs. list_of_dlist (Dlist xs) = xs}"} are invariant,
+  and in our example this is exactly @{term "{xs. distinct xs}"}.
+  
+  The primitive operations on @{typ "'a dlist"} are specified
+  indirectly using the projection @{const list_of_dlist}.  For
+  the empty @{text "dlist"}, @{const Dlist.empty}, we finally want
+  the code equation
+*}
+
+text %quote {*
+  @{term "Dlist.empty = Dlist []"}
+*}
+
+text {*
+  \noindent This we have to prove indirectly as follows:
+*}
+
+lemma %quote [code abstract]:
+  "list_of_dlist Dlist.empty = []"
+  by (fact list_of_dlist_empty)
+
+text {*
+  \noindent This equation logically encodes both the desired code
+  equation and that the expression @{const Dlist} is applied to obeys
+  the implicit invariant.  Equations for insertion and removal are
+  similar:
+*}
+
+lemma %quote [code abstract]:
+  "list_of_dlist (Dlist.insert x dxs) = List.insert x (list_of_dlist dxs)"
+  by (fact list_of_dlist_insert)
+
+lemma %quote [code abstract]:
+  "list_of_dlist (Dlist.remove x dxs) = remove1 x (list_of_dlist dxs)"
+  by (fact list_of_dlist_remove)
+
+text {*
+  \noindent Then the corresponding code is as follows:
+*}
+
+text %quote {*
+  @{code_stmts Dlist.empty Dlist.insert Dlist.remove list_of_dlist (Haskell)}
+*} (*(types) dlist (consts) dempty dinsert dremove list_of List.member insert remove *)
+
+text {*
+  Typical data structures implemented by representations involving
+  invariants are available in the library, e.g.~theories @{theory
+  Fset} and @{theory Mapping} specify sets (type @{typ "'a fset"}) and
+  key-value-mappings (type @{typ "('a, 'b) mapping"}) respectively;
+  these can be implemented by distinct lists as presented here as
+  example (theory @{theory Dlist}) and red-black-trees respectively
+  (theory @{theory RBT}).
 *}
 
 end
--- a/doc-src/Codegen/Thy/Setup.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/Thy/Setup.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -1,14 +1,30 @@
 theory Setup
-imports Complex_Main More_List
+imports Complex_Main More_List RBT Dlist Mapping
 uses
   "../../antiquote_setup.ML"
   "../../more_antiquote.ML"
 begin
 
 ML {* no_document use_thys
-  ["Efficient_Nat", "Code_Char_chr", "Product_ord", "~~/src/HOL/Imperative_HOL/Imperative_HOL",
+  ["Efficient_Nat", "Code_Char_chr", "Product_ord",
+   "~~/src/HOL/Imperative_HOL/Imperative_HOL",
    "~~/src/HOL/Decision_Procs/Ferrack"] *}
 
+setup {*
+let
+  val typ = Simple_Syntax.read_typ;
+  val typeT = Syntax.typeT;
+  val spropT = Syntax.spropT;
+in
+  Sign.del_modesyntax_i (Symbol.xsymbolsN, false) [
+    ("_constrain", typ "logic => type => logic", Mixfix ("_\<Colon>_", [4, 0], 3)),
+    ("_constrain", [spropT, typeT] ---> spropT, Mixfix ("_\<Colon>_", [4, 0], 3))]
+  #> Sign.add_modesyntax_i (Symbol.xsymbolsN, false) [
+      ("_constrain", typ "logic => type => logic", Mixfix ("_ \<Colon>  _", [4, 0], 3)),
+      ("_constrain", [spropT, typeT] ---> spropT, Mixfix ("_ \<Colon> _", [4, 0], 3))]
+end
+*}
+
 setup {* Code_Target.set_default_code_width 74 *}
 
 ML_command {* Unsynchronized.reset unique_names *}
--- a/doc-src/Codegen/Thy/document/Adaptation.tex	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Adaptation.tex	Thu Aug 19 17:41:52 2010 +0200
@@ -302,12 +302,12 @@
 \endisadelimquotett
 %
 \begin{isamarkuptext}%
-\noindent The \hyperlink{command.code-type}{\mbox{\isa{\isacommand{code{\isacharunderscore}type}}}} command takes a type constructor
+\noindent The \indexdef{}{command}{code\_type}\hypertarget{command.code-type}{\hyperlink{command.code-type}{\mbox{\isa{\isacommand{code{\isacharunderscore}type}}}}} command takes a type constructor
   as arguments together with a list of custom serialisations.  Each
   custom serialisation starts with a target language identifier
   followed by an expression, which during code serialisation is
   inserted whenever the type constructor would occur.  For constants,
-  \hyperlink{command.code-const}{\mbox{\isa{\isacommand{code{\isacharunderscore}const}}}} implements the corresponding mechanism.  Each
+  \indexdef{}{command}{code\_const}\hypertarget{command.code-const}{\hyperlink{command.code-const}{\mbox{\isa{\isacommand{code{\isacharunderscore}const}}}}} implements the corresponding mechanism.  Each
   ``\verb|_|'' in a serialisation expression is treated as a
   placeholder for the type constructor's (the constant's) arguments.%
 \end{isamarkuptext}%
@@ -416,7 +416,7 @@
   avoided to be used for new declarations.  Initially, this table
   typically contains the keywords of the target language.  It can be
   extended manually, thus avoiding accidental overwrites, using the
-  \hyperlink{command.code-reserved}{\mbox{\isa{\isacommand{code{\isacharunderscore}reserved}}}} command:%
+  \indexdef{}{command}{code\_reserved}\hypertarget{command.code-reserved}{\hyperlink{command.code-reserved}{\mbox{\isa{\isacommand{code{\isacharunderscore}reserved}}}}} command:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -501,7 +501,7 @@
 For convenience, the default \isa{HOL} setup for \isa{Haskell}
   maps the \isa{eq} class to its counterpart in \isa{Haskell},
   giving custom serialisations for the class \isa{eq} (by command
-  \hyperlink{command.code-class}{\mbox{\isa{\isacommand{code{\isacharunderscore}class}}}}) and its operation \isa{eq{\isacharunderscore}class{\isachardot}eq}%
+  \indexdef{}{command}{code\_class}\hypertarget{command.code-class}{\hyperlink{command.code-class}{\mbox{\isa{\isacommand{code{\isacharunderscore}class}}}}}) and its operation \isa{eq{\isacharunderscore}class{\isachardot}eq}%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -577,7 +577,7 @@
 \begin{isamarkuptext}%
 \noindent The code generator would produce an additional instance,
   which of course is rejected by the \isa{Haskell} compiler.  To
-  suppress this additional instance, use \isa{code{\isacharunderscore}instance}:%
+  suppress this additional instance, use \indexdef{}{command}{code\_instance}\hypertarget{command.code-instance}{\hyperlink{command.code-instance}{\mbox{\isa{\isacommand{code{\isacharunderscore}instance}}}}}:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -602,7 +602,7 @@
 %
 \begin{isamarkuptext}%
 In rare cases it is necessary to \emph{enrich} the context of a
-  target language; this is accomplished using the \hyperlink{command.code-include}{\mbox{\isa{\isacommand{code{\isacharunderscore}include}}}} command:%
+  target language; this is accomplished using the \indexdef{}{command}{code\_include}\hypertarget{command.code-include}{\hyperlink{command.code-include}{\mbox{\isa{\isacommand{code{\isacharunderscore}include}}}}} command:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/Codegen/Thy/document/Foundations.tex	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Foundations.tex	Thu Aug 19 17:41:52 2010 +0200
@@ -182,9 +182,9 @@
   \secref{eff_nat}) uses this interface.
 
   \noindent The current setup of the preprocessor may be inspected
-  using the \hyperlink{command.print-codeproc}{\mbox{\isa{\isacommand{print{\isacharunderscore}codeproc}}}} command.  \hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}}
-  (see \secref{sec:equations}) provides a convenient mechanism to
-  inspect the impact of a preprocessor setup on code equations.
+  using the \indexdef{}{command}{print\_codeproc}\hypertarget{command.print-codeproc}{\hyperlink{command.print-codeproc}{\mbox{\isa{\isacommand{print{\isacharunderscore}codeproc}}}}} command.  \indexdef{}{command}{code\_thms}\hypertarget{command.code-thms}{\hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}}} (see \secref{sec:equations}) provides a convenient
+  mechanism to inspect the impact of a preprocessor setup on code
+  equations.
 
   \begin{warn}
     Attribute \hyperlink{attribute.code-unfold}{\mbox{\isa{code{\isacharunderscore}unfold}}} also applies to the
@@ -275,8 +275,7 @@
   equations (before preprocessing) and code equations (after
   preprocessing).
 
-  The first can be listed (among other data) using the \hyperlink{command.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}} command (later on in \secref{sec:refinement}
-  it will be shown how these code equations can be changed).
+  The first can be listed (among other data) using the \indexdef{}{command}{print\_codesetup}\hypertarget{command.print-codesetup}{\hyperlink{command.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}}} command.
 
   The code equations after preprocessing are already are blueprint of
   the generated program and can be inspected using the \hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}} command:%
@@ -300,7 +299,7 @@
 \begin{isamarkuptext}%
 \noindent This prints a table with the code equations for \isa{dequeue}, including \emph{all} code equations those equations depend
   on recursively.  These dependencies themselves can be visualized using
-  the \hyperlink{command.code-deps}{\mbox{\isa{\isacommand{code{\isacharunderscore}deps}}}} command.%
+  the \indexdef{}{command}{code\_deps}\hypertarget{command.code-deps}{\hyperlink{command.code-deps}{\mbox{\isa{\isacommand{code{\isacharunderscore}deps}}}}} command.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -503,7 +502,7 @@
   of as function definitions which always fail,
   since there is never a successful pattern match on the left hand
   side.  In order to categorise a constant into that category
-  explicitly, use \hyperlink{command.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}}:%
+  explicitly, use \indexdef{}{command}{code\_abort}\hypertarget{command.code-abort}{\hyperlink{command.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}}}:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/Codegen/Thy/document/Further.tex	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Further.tex	Thu Aug 19 17:41:52 2010 +0200
@@ -145,7 +145,7 @@
 \endisadelimquote
 %
 \begin{isamarkuptext}%
-After an interpretation of this locale (say, \hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}} \isa{fun{\isacharunderscore}power{\isacharcolon}} \isa{{\isachardoublequote}power\ {\isacharparenleft}{\isasymlambda}n\ {\isacharparenleft}f\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}{\isachardot}\ f\ {\isacharcircum}{\isacharcircum}\ n{\isacharparenright}{\isachardoublequote}}), one would expect to have a constant \isa{fun{\isacharunderscore}power{\isachardot}powers\ {\isacharcolon}{\isacharcolon}\ nat\ list\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} for which code
+After an interpretation of this locale (say, \indexdef{}{command}{interpretation}\hypertarget{command.interpretation}{\hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}}} \isa{fun{\isacharunderscore}power{\isacharcolon}} \isa{{\isachardoublequote}power\ {\isacharparenleft}{\isasymlambda}n\ {\isacharparenleft}f\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}{\isachardot}\ f\ {\isacharcircum}{\isacharcircum}\ n{\isacharparenright}{\isachardoublequote}}), one would expect to have a constant \isa{fun{\isacharunderscore}power{\isachardot}powers\ {\isacharcolon}{\isacharcolon}\ nat\ list\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} for which code
   can be generated.  But this not the case: internally, the term
   \isa{fun{\isacharunderscore}power{\isachardot}powers} is an abbreviation for the foundational
   term \isa{{\isachardoublequote}power{\isachardot}powers\ {\isacharparenleft}{\isasymlambda}n\ {\isacharparenleft}f\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}{\isachardot}\ f\ {\isacharcircum}{\isacharcircum}\ n{\isacharparenright}{\isachardoublequote}}
@@ -246,156 +246,15 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
-\isamarkupsubsection{Evaluation oracle%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Code generation may also be used to \emph{evaluate} expressions
-  (using \isa{SML} as target language of course).
-  For instance, the \hyperlink{command.value}{\mbox{\isa{\isacommand{value}}}} reduces an expression to a
-  normal form with respect to the underlying code equations:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{value}\isamarkupfalse%
-\ {\isachardoublequoteopen}{\isadigit{4}}{\isadigit{2}}\ {\isacharslash}\ {\isacharparenleft}{\isadigit{1}}{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ rat{\isacharparenright}{\isachardoublequoteclose}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent will display \isa{{\isadigit{7}}\ {\isacharslash}\ {\isadigit{2}}}.
-
-  The \hyperlink{method.eval}{\mbox{\isa{eval}}} method tries to reduce a goal by code generation to \isa{True}
-  and solves it in that case, but fails otherwise:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{lemma}\isamarkupfalse%
-\ {\isachardoublequoteopen}{\isadigit{4}}{\isadigit{2}}\ {\isacharslash}\ {\isacharparenleft}{\isadigit{1}}{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ rat{\isacharparenright}\ {\isacharequal}\ {\isadigit{7}}\ {\isacharslash}\ {\isadigit{2}}{\isachardoublequoteclose}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ eval%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent The soundness of the \hyperlink{method.eval}{\mbox{\isa{eval}}} method depends crucially 
-  on the correctness of the code generator;  this is one of the reasons
-  why you should not use adaptation (see \secref{sec:adaptation}) frivolously.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Building evaluators%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-FIXME%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsubsection{Code antiquotation%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-In scenarios involving techniques like reflection it is quite common
-  that code generated from a theory forms the basis for implementing
-  a proof procedure in \isa{SML}.  To facilitate interfacing of generated code
-  with system code, the code generator provides a \isa{code} antiquotation:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{datatype}\isamarkupfalse%
-\ form\ {\isacharequal}\ T\ {\isacharbar}\ F\ {\isacharbar}\ And\ form\ form\ {\isacharbar}\ Or\ form\ form\ %
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isadelimquotett
-\ %
-\endisadelimquotett
-%
-\isatagquotett
-\isacommand{ML}\isamarkupfalse%
-\ {\isacharverbatimopen}\isanewline
-\ \ fun\ eval{\isacharunderscore}form\ %
-\isaantiq
-code\ T%
-\endisaantiq
-\ {\isacharequal}\ true\isanewline
-\ \ \ \ {\isacharbar}\ eval{\isacharunderscore}form\ %
-\isaantiq
-code\ F%
-\endisaantiq
-\ {\isacharequal}\ false\isanewline
-\ \ \ \ {\isacharbar}\ eval{\isacharunderscore}form\ {\isacharparenleft}%
-\isaantiq
-code\ And%
-\endisaantiq
-\ {\isacharparenleft}p{\isacharcomma}\ q{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline
-\ \ \ \ \ \ \ \ eval{\isacharunderscore}form\ p\ andalso\ eval{\isacharunderscore}form\ q\isanewline
-\ \ \ \ {\isacharbar}\ eval{\isacharunderscore}form\ {\isacharparenleft}%
-\isaantiq
-code\ Or%
-\endisaantiq
-\ {\isacharparenleft}p{\isacharcomma}\ q{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline
-\ \ \ \ \ \ \ \ eval{\isacharunderscore}form\ p\ orelse\ eval{\isacharunderscore}form\ q{\isacharsemicolon}\isanewline
-{\isacharverbatimclose}%
-\endisatagquotett
-{\isafoldquotett}%
-%
-\isadelimquotett
-%
-\endisadelimquotett
-%
-\begin{isamarkuptext}%
-\noindent \isa{code} takes as argument the name of a constant;  after the
-  whole \isa{SML} is read, the necessary code is generated transparently
-  and the corresponding constant names are inserted.  This technique also
-  allows to use pattern matching on constructors stemming from compiled
-  \isa{datatypes}.
-
-  For a less simplistic example, theory \hyperlink{theory.Ferrack}{\mbox{\isa{Ferrack}}} is
-  a good reference.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
 \isamarkupsubsection{ML system interfaces \label{sec:ml}%
 }
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-Since the code generator framework not only aims to provide
-  a nice Isar interface but also to form a base for
-  code-generation-based applications, here a short
-  description of the most important ML interfaces.%
+Since the code generator framework not only aims to provide a nice
+  Isar interface but also to form a base for code-generation-based
+  applications, here a short description of the most fundamental ML
+  interfaces.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -411,12 +270,14 @@
 %
 \begin{isamarkuptext}%
 \begin{mldecls}
+  \indexdef{}{ML}{Code.read\_const}\verb|Code.read_const: theory -> string -> string| \\
   \indexdef{}{ML}{Code.add\_eqn}\verb|Code.add_eqn: thm -> theory -> theory| \\
   \indexdef{}{ML}{Code.del\_eqn}\verb|Code.del_eqn: thm -> theory -> theory| \\
   \indexdef{}{ML}{Code\_Preproc.map\_pre}\verb|Code_Preproc.map_pre: (simpset -> simpset) -> theory -> theory| \\
   \indexdef{}{ML}{Code\_Preproc.map\_post}\verb|Code_Preproc.map_post: (simpset -> simpset) -> theory -> theory| \\
-  \indexdef{}{ML}{Code\_Preproc.add\_functrans}\verb|Code_Preproc.add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option)|\isasep\isanewline%
-\verb|    -> theory -> theory| \\
+  \indexdef{}{ML}{Code\_Preproc.add\_functrans}\verb|Code_Preproc.add_functrans: |\isasep\isanewline%
+\verb|    string * (theory -> (thm * bool) list -> (thm * bool) list option)|\isasep\isanewline%
+\verb|      -> theory -> theory| \\
   \indexdef{}{ML}{Code\_Preproc.del\_functrans}\verb|Code_Preproc.del_functrans: string -> theory -> theory| \\
   \indexdef{}{ML}{Code.add\_datatype}\verb|Code.add_datatype: (string * typ) list -> theory -> theory| \\
   \indexdef{}{ML}{Code.get\_type}\verb|Code.get_type: theory -> string|\isasep\isanewline%
@@ -426,6 +287,9 @@
 
   \begin{description}
 
+  \item \verb|Code.read_const|~\isa{thy}~\isa{s}
+     reads a constant as a concrete term expression \isa{s}.
+
   \item \verb|Code.add_eqn|~\isa{thm}~\isa{thy} adds function
      theorem \isa{thm} to executable content.
 
@@ -466,37 +330,6 @@
 %
 \endisadelimmlref
 %
-\isamarkupsubsubsection{Auxiliary%
-}
-\isamarkuptrue%
-%
-\isadelimmlref
-%
-\endisadelimmlref
-%
-\isatagmlref
-%
-\begin{isamarkuptext}%
-\begin{mldecls}
-  \indexdef{}{ML}{Code.read\_const}\verb|Code.read_const: theory -> string -> string|
-  \end{mldecls}
-
-  \begin{description}
-
-  \item \verb|Code.read_const|~\isa{thy}~\isa{s}
-     reads a constant as a concrete term expression \isa{s}.
-
-  \end{description}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagmlref
-{\isafoldmlref}%
-%
-\isadelimmlref
-%
-\endisadelimmlref
-%
 \isamarkupsubsubsection{Data depending on the theory's executable content%
 }
 \isamarkuptrue%
--- a/doc-src/Codegen/Thy/document/Inductive_Predicate.tex	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Inductive_Predicate.tex	Thu Aug 19 17:41:52 2010 +0200
@@ -37,7 +37,7 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-The \isa{predicate{\isacharunderscore}compiler} is an extension of the code generator
+The \isa{predicate\ compiler} is an extension of the code generator
   which turns inductive specifications into equational ones, from
   which in turn executable code can be generated.  The mechanisms of
   this compiler are described in detail in
--- a/doc-src/Codegen/Thy/document/Introduction.tex	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Introduction.tex	Thu Aug 19 17:41:52 2010 +0200
@@ -55,10 +55,11 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-In a HOL theory, the \hyperlink{command.datatype}{\mbox{\isa{\isacommand{datatype}}}} and \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}/\hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}}/\hyperlink{command.fun}{\mbox{\isa{\isacommand{fun}}}} declarations form the
-  core of a functional programming language.  By default equational
-  theorems stemming from those are used for generated code, therefore
-  \qt{naive} code generation can proceed without further ado.
+In a HOL theory, the \indexdef{}{command}{datatype}\hypertarget{command.datatype}{\hyperlink{command.datatype}{\mbox{\isa{\isacommand{datatype}}}}} and \indexdef{}{command}{definition}\hypertarget{command.definition}{\hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}}/\indexdef{}{command}{primrec}\hypertarget{command.primrec}{\hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}}}/\indexdef{}{command}{fun}\hypertarget{command.fun}{\hyperlink{command.fun}{\mbox{\isa{\isacommand{fun}}}}} declarations
+  form the core of a functional programming language.  By default
+  equational theorems stemming from those are used for generated code,
+  therefore \qt{naive} code generation can proceed without further
+  ado.
 
   For example, here a simple \qt{implementation} of amortised queues:%
 \end{isamarkuptext}%
@@ -184,12 +185,12 @@
 \endisadelimquote
 %
 \begin{isamarkuptext}%
-\noindent The \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} command takes a space-separated
-  list of constants for which code shall be generated; anything else
-  needed for those is added implicitly.  Then follows a 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
+\noindent The \indexdef{}{command}{export\_code}\hypertarget{command.export-code}{\hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}}} command takes a
+  space-separated list of constants for which code shall be generated;
+  anything else needed for those is added implicitly.  Then follows a
+  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
   module name (with extension \isa{{\isachardot}hs}) is written:%
 \end{isamarkuptext}%
--- a/doc-src/Codegen/Thy/document/Refinement.tex	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Refinement.tex	Thu Aug 19 17:41:52 2010 +0200
@@ -268,7 +268,7 @@
   constructor must be of type \isa{{\isasymtau}\ {\isacharequal}\ {\isasymdots}\ {\isasymRightarrow}\ {\isasymkappa}\ {\isasymalpha}\isactrlisub {\isadigit{1}}\ {\isasymdots}\ {\isasymalpha}\isactrlisub n} where \isa{{\isacharbraceleft}{\isasymalpha}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlisub n{\isacharbraceright}} is exactly the set of \emph{all} type variables in
   \isa{{\isasymtau}}; then \isa{{\isasymkappa}} is its corresponding datatype.  The
   HOL datatype package by default registers any new datatype with its
-  constructors, but this may be changed using \hyperlink{command.code-datatype}{\mbox{\isa{\isacommand{code{\isacharunderscore}datatype}}}}; the currently chosen constructors can be inspected
+  constructors, but this may be changed using \indexdef{}{command}{code\_datatype}\hypertarget{command.code-datatype}{\hyperlink{command.code-datatype}{\mbox{\isa{\isacommand{code{\isacharunderscore}datatype}}}}}; the currently chosen constructors can be inspected
   using the \hyperlink{command.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}} command.
 
   Equipped with this, we are able to prove the following equations
@@ -396,7 +396,7 @@
 \begin{isamarkuptext}%
 The same techniques can also be applied to types which are not
   specified as datatypes, e.g.~type \isa{int} is originally specified
-  as quotient type by means of \hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}}, but for code
+  as quotient type by means of \indexdef{}{command}{typedef}\hypertarget{command.typedef}{\hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}}}, but for code
   generation constants allowing construction of binary numeral values
   are used as constructors for \isa{int}.
 
@@ -410,7 +410,227 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-FIXME%
+Datatype representation involving invariants require a dedicated
+  setup for the type and its primitive operations.  As a running
+  example, we implement a type \isa{{\isacharprime}a\ dlist} of list consisting
+  of distinct elements.
+
+  The first step is to decide on which representation the abstract
+  type (in our example \isa{{\isacharprime}a\ dlist}) should be implemented.
+  Here we choose \isa{{\isacharprime}a\ list}.  Then a conversion from the concrete
+  type to the abstract type must be specified, here:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\isatagquote
+%
+\begin{isamarkuptext}%
+\isa{Dlist\ {\isasymColon}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ dlist}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagquote
+{\isafoldquote}%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\begin{isamarkuptext}%
+\noindent Next follows the specification of a suitable \emph{projection},
+  i.e.~a conversion from abstract to concrete type:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\isatagquote
+%
+\begin{isamarkuptext}%
+\isa{list{\isacharunderscore}of{\isacharunderscore}dlist\ {\isasymColon}\ {\isacharprime}a\ dlist\ {\isasymRightarrow}\ {\isacharprime}a\ list}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagquote
+{\isafoldquote}%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\begin{isamarkuptext}%
+\noindent This projection must be specified such that the following
+  \emph{abstract datatype certificate} can be proven:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\isatagquote
+\isacommand{lemma}\isamarkupfalse%
+\ {\isacharbrackleft}code\ abstype{\isacharbrackright}{\isacharcolon}\isanewline
+\ \ {\isachardoublequoteopen}Dlist\ {\isacharparenleft}list{\isacharunderscore}of{\isacharunderscore}dlist\ dxs{\isacharparenright}\ {\isacharequal}\ dxs{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}fact\ Dlist{\isacharunderscore}list{\isacharunderscore}of{\isacharunderscore}dlist{\isacharparenright}%
+\endisatagquote
+{\isafoldquote}%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\begin{isamarkuptext}%
+\noindent Note that so far the invariant on representations
+  (\isa{distinct\ {\isasymColon}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ bool}) has never been mentioned explicitly:
+  the invariant is only referred to implicitly: all values in
+  set \isa{{\isacharbraceleft}xs{\isachardot}\ list{\isacharunderscore}of{\isacharunderscore}dlist\ {\isacharparenleft}Dlist\ xs{\isacharparenright}\ {\isacharequal}\ xs{\isacharbraceright}} are invariant,
+  and in our example this is exactly \isa{{\isacharbraceleft}xs{\isachardot}\ distinct\ xs{\isacharbraceright}}.
+  
+  The primitive operations on \isa{{\isacharprime}a\ dlist} are specified
+  indirectly using the projection \isa{list{\isacharunderscore}of{\isacharunderscore}dlist}.  For
+  the empty \isa{dlist}, \isa{Dlist{\isachardot}empty}, we finally want
+  the code equation%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\isatagquote
+%
+\begin{isamarkuptext}%
+\isa{Dlist{\isachardot}empty\ {\isacharequal}\ Dlist\ {\isacharbrackleft}{\isacharbrackright}}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagquote
+{\isafoldquote}%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\begin{isamarkuptext}%
+\noindent This we have to prove indirectly as follows:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\isatagquote
+\isacommand{lemma}\isamarkupfalse%
+\ {\isacharbrackleft}code\ abstract{\isacharbrackright}{\isacharcolon}\isanewline
+\ \ {\isachardoublequoteopen}list{\isacharunderscore}of{\isacharunderscore}dlist\ Dlist{\isachardot}empty\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}fact\ list{\isacharunderscore}of{\isacharunderscore}dlist{\isacharunderscore}empty{\isacharparenright}%
+\endisatagquote
+{\isafoldquote}%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\begin{isamarkuptext}%
+\noindent This equation logically encodes both the desired code
+  equation and that the expression \isa{Dlist} is applied to obeys
+  the implicit invariant.  Equations for insertion and removal are
+  similar:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\isatagquote
+\isacommand{lemma}\isamarkupfalse%
+\ {\isacharbrackleft}code\ abstract{\isacharbrackright}{\isacharcolon}\isanewline
+\ \ {\isachardoublequoteopen}list{\isacharunderscore}of{\isacharunderscore}dlist\ {\isacharparenleft}Dlist{\isachardot}insert\ x\ dxs{\isacharparenright}\ {\isacharequal}\ List{\isachardot}insert\ x\ {\isacharparenleft}list{\isacharunderscore}of{\isacharunderscore}dlist\ dxs{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}fact\ list{\isacharunderscore}of{\isacharunderscore}dlist{\isacharunderscore}insert{\isacharparenright}\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ {\isacharbrackleft}code\ abstract{\isacharbrackright}{\isacharcolon}\isanewline
+\ \ {\isachardoublequoteopen}list{\isacharunderscore}of{\isacharunderscore}dlist\ {\isacharparenleft}Dlist{\isachardot}remove\ x\ dxs{\isacharparenright}\ {\isacharequal}\ remove{\isadigit{1}}\ x\ {\isacharparenleft}list{\isacharunderscore}of{\isacharunderscore}dlist\ dxs{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}fact\ list{\isacharunderscore}of{\isacharunderscore}dlist{\isacharunderscore}remove{\isacharparenright}%
+\endisatagquote
+{\isafoldquote}%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\begin{isamarkuptext}%
+\noindent Then the corresponding code is as follows:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\isatagquote
+%
+\begin{isamarkuptext}%
+\isatypewriter%
+\noindent%
+\hspace*{0pt}module Example where {\char123}\\
+\hspace*{0pt}\\
+\hspace*{0pt}newtype Dlist a = Dlist [a];\\
+\hspace*{0pt}\\
+\hspace*{0pt}empty ::~forall a.~Dlist a;\\
+\hspace*{0pt}empty = Dlist [];\\
+\hspace*{0pt}\\
+\hspace*{0pt}member ::~forall a.~(Eq a) => [a] -> a -> Bool;\\
+\hspace*{0pt}member [] y = False;\\
+\hspace*{0pt}member (x :~xs) y = x == y || member xs y;\\
+\hspace*{0pt}\\
+\hspace*{0pt}inserta ::~forall a.~(Eq a) => a -> [a] -> [a];\\
+\hspace*{0pt}inserta x xs = (if member xs x then xs else x :~xs);\\
+\hspace*{0pt}\\
+\hspace*{0pt}list{\char95}of{\char95}dlist ::~forall a.~Dlist a -> [a];\\
+\hspace*{0pt}list{\char95}of{\char95}dlist (Dlist x) = x;\\
+\hspace*{0pt}\\
+\hspace*{0pt}insert ::~forall a.~(Eq a) => a -> Dlist a -> Dlist a;\\
+\hspace*{0pt}insert x dxs = Dlist (inserta x (list{\char95}of{\char95}dlist dxs));\\
+\hspace*{0pt}\\
+\hspace*{0pt}remove1 ::~forall a.~(Eq a) => a -> [a] -> [a];\\
+\hspace*{0pt}remove1 x [] = [];\\
+\hspace*{0pt}remove1 x (y :~xs) = (if x == y then xs else y :~remove1 x xs);\\
+\hspace*{0pt}\\
+\hspace*{0pt}remove ::~forall a.~(Eq a) => a -> Dlist a -> Dlist a;\\
+\hspace*{0pt}remove x dxs = Dlist (remove1 x (list{\char95}of{\char95}dlist dxs));\\
+\hspace*{0pt}\\
+\hspace*{0pt}{\char125}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagquote
+{\isafoldquote}%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\begin{isamarkuptext}%
+Typical data structures implemented by representations involving
+  invariants are available in the library, e.g.~theories \hyperlink{theory.Fset}{\mbox{\isa{Fset}}} and \hyperlink{theory.Mapping}{\mbox{\isa{Mapping}}} specify sets (type \isa{{\isacharprime}a\ fset}) and
+  key-value-mappings (type \isa{{\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ mapping}) respectively;
+  these can be implemented by distinct lists as presented here as
+  example (theory \hyperlink{theory.Dlist}{\mbox{\isa{Dlist}}}) and red-black-trees respectively
+  (theory \hyperlink{theory.RBT}{\mbox{\isa{RBT}}}).%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/Codegen/codegen.tex	Thu Aug 19 17:40:44 2010 +0200
+++ b/doc-src/Codegen/codegen.tex	Thu Aug 19 17:41:52 2010 +0200
@@ -35,6 +35,7 @@
 \input{Thy/document/Refinement.tex}
 \input{Thy/document/Inductive_Predicate.tex}
 \input{Thy/document/Adaptation.tex}
+\input{Thy/document/Evaluation.tex}
 \input{Thy/document/Further.tex}
 
 \begingroup
--- a/src/CCL/Set.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/CCL/Set.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -4,8 +4,6 @@
 imports FOL
 begin
 
-global
-
 typedecl 'a set
 arities set :: ("term") "term"
 
@@ -46,8 +44,6 @@
   "ALL x:A. P"  == "CONST Ball(A, %x. P)"
   "EX x:A. P"   == "CONST Bex(A, %x. P)"
 
-local
-
 axioms
   mem_Collect_iff:       "(a : {x. P(x)}) <-> P(a)"
   set_extension:         "A=B <-> (ALL x. x:A <-> x:B)"
--- a/src/CCL/Wfd.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/CCL/Wfd.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -423,13 +423,13 @@
   @{thms canTs} @ @{thms icanTs} @ @{thms applyT2} @ @{thms ncanTs} @ @{thms incanTs} @
   @{thms precTs} @ @{thms letrecTs} @ @{thms letT} @ @{thms Subtype_canTs};
 
-fun bvars (Const("all",_) $ Abs(s,_,t)) l = bvars t (s::l)
+fun bvars (Const(@{const_name all},_) $ Abs(s,_,t)) l = bvars t (s::l)
   | bvars _ l = l
 
-fun get_bno l n (Const("all",_) $ Abs(s,_,t)) = get_bno (s::l) n t
-  | get_bno l n (Const("Trueprop",_) $ t) = get_bno l n t
-  | get_bno l n (Const("Ball",_) $ _ $ Abs(s,_,t)) = get_bno (s::l) (n+1) t
-  | get_bno l n (Const("mem",_) $ t $ _) = get_bno l n t
+fun get_bno l n (Const(@{const_name all},_) $ Abs(s,_,t)) = get_bno (s::l) n t
+  | get_bno l n (Const(@{const_name Trueprop},_) $ t) = get_bno l n t
+  | get_bno l n (Const(@{const_name Ball},_) $ _ $ Abs(s,_,t)) = get_bno (s::l) (n+1) t
+  | get_bno l n (Const(@{const_name mem},_) $ t $ _) = get_bno l n t
   | get_bno l n (t $ s) = get_bno l n t
   | get_bno l n (Bound m) = (m-length(l),n)
 
@@ -450,7 +450,7 @@
 
 fun is_rigid_prog t =
      case (Logic.strip_assums_concl t) of
-        (Const("Trueprop",_) $ (Const("mem",_) $ a $ _)) => null (Term.add_vars a [])
+        (Const(@{const_name Trueprop},_) $ (Const(@{const_name mem},_) $ a $ _)) => null (Term.add_vars a [])
        | _ => false
 in
 
--- a/src/FOL/IFOL.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/FOL/IFOL.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -28,8 +28,6 @@
 
 setup PureThy.old_appl_syntax_setup
 
-global
-
 classes "term"
 default_sort "term"
 
@@ -87,8 +85,6 @@
   Ex        (binder "\<exists>" 10) and
   Ex1       (binder "\<exists>!" 10)
 
-local
-
 finalconsts
   False All Ex
   "op ="
--- a/src/FOL/fologic.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/FOL/fologic.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -37,48 +37,48 @@
 structure FOLogic: FOLOGIC =
 struct
 
-val oT = Type("o",[]);
+val oT = Type(@{type_name o},[]);
 
-val Trueprop = Const("Trueprop", oT-->propT);
+val Trueprop = Const(@{const_name Trueprop}, oT-->propT);
 
 fun mk_Trueprop P = Trueprop $ P;
 
-fun dest_Trueprop (Const ("Trueprop", _) $ P) = P
+fun dest_Trueprop (Const (@{const_name Trueprop}, _) $ P) = P
   | dest_Trueprop t = raise TERM ("dest_Trueprop", [t]);
 
 
 (* Logical constants *)
 
-val not = Const ("Not", oT --> oT);
-val conj = Const("op &", [oT,oT]--->oT);
-val disj = Const("op |", [oT,oT]--->oT);
-val imp = Const("op -->", [oT,oT]--->oT)
-val iff = Const("op <->", [oT,oT]--->oT);
+val not = Const (@{const_name Not}, oT --> oT);
+val conj = Const(@{const_name "op &"}, [oT,oT]--->oT);
+val disj = Const(@{const_name "op |"}, [oT,oT]--->oT);
+val imp = Const(@{const_name "op -->"}, [oT,oT]--->oT)
+val iff = Const(@{const_name "op <->"}, [oT,oT]--->oT);
 
 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_iff (t1, t2) = iff $ t1 $ t2;
 
-fun dest_imp (Const("op -->",_) $ A $ B) = (A, B)
+fun dest_imp (Const(@{const_name "op -->"},_) $ A $ B) = (A, B)
   | dest_imp  t = raise TERM ("dest_imp", [t]);
 
-fun dest_conj (Const ("op &", _) $ t $ t') = t :: dest_conj t'
+fun dest_conj (Const (@{const_name "op &"}, _) $ t $ t') = t :: dest_conj t'
   | dest_conj t = [t];
 
-fun dest_iff (Const("op <->",_) $ A $ B) = (A, B)
+fun dest_iff (Const(@{const_name "op <->"},_) $ A $ B) = (A, B)
   | dest_iff  t = raise TERM ("dest_iff", [t]);
 
-fun eq_const T = Const ("op =", [T, T] ---> oT);
+fun eq_const T = Const (@{const_name "op ="}, [T, T] ---> oT);
 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 (@{const_name "op ="}, _) $ lhs $ rhs) = (lhs, rhs)
   | dest_eq t = raise TERM ("dest_eq", [t])
 
-fun all_const T = Const ("All", [T --> oT] ---> oT);
+fun all_const T = Const (@{const_name All}, [T --> oT] ---> oT);
 fun mk_all (Free(x,T),P) = all_const T $ (absfree (x,T,P));
 
-fun exists_const T = Const ("Ex", [T --> oT] ---> oT);
+fun exists_const T = Const (@{const_name Ex}, [T --> oT] ---> oT);
 fun mk_exists (Free(x,T),P) = exists_const T $ (absfree (x,T,P));
 
 
--- a/src/FOL/intprover.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/FOL/intprover.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -40,22 +40,22 @@
   step of an intuitionistic proof.
 *)
 val safe_brls = sort (make_ord lessb)
-    [ (true, thm "FalseE"), (false, thm "TrueI"), (false, thm "refl"),
-      (false, thm "impI"), (false, thm "notI"), (false, thm "allI"),
-      (true, thm "conjE"), (true, thm "exE"),
-      (false, thm "conjI"), (true, thm "conj_impE"),
-      (true, thm "disj_impE"), (true, thm "disjE"), 
-      (false, thm "iffI"), (true, thm "iffE"), (true, thm "not_to_imp") ];
+    [ (true, @{thm FalseE}), (false, @{thm TrueI}), (false, @{thm refl}),
+      (false, @{thm impI}), (false, @{thm notI}), (false, @{thm allI}),
+      (true, @{thm conjE}), (true, @{thm exE}),
+      (false, @{thm conjI}), (true, @{thm conj_impE}),
+      (true, @{thm disj_impE}), (true, @{thm disjE}), 
+      (false, @{thm iffI}), (true, @{thm iffE}), (true, @{thm not_to_imp}) ];
 
 val haz_brls =
-    [ (false, thm "disjI1"), (false, thm "disjI2"), (false, thm "exI"), 
-      (true, thm "allE"), (true, thm "not_impE"), (true, thm "imp_impE"), (true, thm "iff_impE"),
-      (true, thm "all_impE"), (true, thm "ex_impE"), (true, thm "impE") ];
+    [ (false, @{thm disjI1}), (false, @{thm disjI2}), (false, @{thm exI}), 
+      (true, @{thm allE}), (true, @{thm not_impE}), (true, @{thm imp_impE}), (true, @{thm iff_impE}),
+      (true, @{thm all_impE}), (true, @{thm ex_impE}), (true, @{thm impE}) ];
 
 val haz_dup_brls =
-    [ (false, thm "disjI1"), (false, thm "disjI2"), (false, thm "exI"),
-      (true, thm "all_dupE"), (true, thm "not_impE"), (true, thm "imp_impE"), (true, thm "iff_impE"),
-      (true, thm "all_impE"), (true, thm "ex_impE"), (true, thm "impE") ];
+    [ (false, @{thm disjI1}), (false, @{thm disjI2}), (false, @{thm exI}),
+      (true, @{thm all_dupE}), (true, @{thm not_impE}), (true, @{thm imp_impE}), (true, @{thm iff_impE}),
+      (true, @{thm all_impE}), (true, @{thm ex_impE}), (true, @{thm impE}) ];
 
 (*0 subgoals vs 1 or more: the p in safep is for positive*)
 val (safe0_brls, safep_brls) =
--- a/src/FOL/simpdata.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/FOL/simpdata.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -8,16 +8,16 @@
 (*Make meta-equalities.  The operator below is Trueprop*)
 
 fun mk_meta_eq th = case concl_of th of
-    _ $ (Const("op =",_)$_$_)   => th RS @{thm eq_reflection}
-  | _ $ (Const("op <->",_)$_$_) => th RS @{thm iff_reflection}
+    _ $ (Const(@{const_name "op ="},_)$_$_)   => th RS @{thm eq_reflection}
+  | _ $ (Const(@{const_name "op <->"},_)$_$_) => th RS @{thm iff_reflection}
   | _                           =>
   error("conclusion must be a =-equality or <->");;
 
 fun mk_eq th = case concl_of th of
     Const("==",_)$_$_           => th
-  | _ $ (Const("op =",_)$_$_)   => mk_meta_eq th
-  | _ $ (Const("op <->",_)$_$_) => mk_meta_eq th
-  | _ $ (Const("Not",_)$_)      => th RS @{thm iff_reflection_F}
+  | _ $ (Const(@{const_name "op ="},_)$_$_)   => mk_meta_eq th
+  | _ $ (Const(@{const_name "op <->"},_)$_$_) => mk_meta_eq th
+  | _ $ (Const(@{const_name Not},_)$_)      => th RS @{thm iff_reflection_F}
   | _                           => th RS @{thm iff_reflection_T};
 
 (*Replace premises x=y, X<->Y by X==Y*)
@@ -32,13 +32,13 @@
       error("Premises and conclusion of congruence rules must use =-equality or <->");
 
 val mksimps_pairs =
-  [("op -->", [@{thm mp}]), ("op &", [@{thm conjunct1}, @{thm conjunct2}]),
-   ("All", [@{thm spec}]), ("True", []), ("False", [])];
+  [(@{const_name "op -->"}, [@{thm mp}]), (@{const_name "op &"}, [@{thm conjunct1}, @{thm conjunct2}]),
+   (@{const_name All}, [@{thm spec}]), (@{const_name True}, []), (@{const_name False}, [])];
 
 fun mk_atomize pairs =
   let fun atoms th =
         (case concl_of th of
-           Const("Trueprop",_) $ p =>
+           Const(@{const_name Trueprop},_) $ p =>
              (case head_of p of
                 Const(a,_) =>
                   (case AList.lookup (op =) pairs a of
@@ -55,11 +55,11 @@
 structure Quantifier1 = Quantifier1Fun(
 struct
   (*abstract syntax*)
-  fun dest_eq((c as Const("op =",_)) $ s $ t) = SOME(c,s,t)
+  fun dest_eq((c as Const(@{const_name "op ="},_)) $ s $ t) = SOME(c,s,t)
     | dest_eq _ = NONE;
-  fun dest_conj((c as Const("op &",_)) $ s $ t) = SOME(c,s,t)
+  fun dest_conj((c as Const(@{const_name "op &"},_)) $ s $ t) = SOME(c,s,t)
     | dest_conj _ = NONE;
-  fun dest_imp((c as Const("op -->",_)) $ s $ t) = SOME(c,s,t)
+  fun dest_imp((c as Const(@{const_name "op -->"},_)) $ s $ t) = SOME(c,s,t)
     | dest_imp _ = NONE;
   val conj = FOLogic.conj
   val imp  = FOLogic.imp
--- a/src/FOLP/IFOLP.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/FOLP/IFOLP.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -12,8 +12,6 @@
 
 setup PureThy.old_appl_syntax_setup
 
-global
-
 classes "term"
 default_sort "term"
 
@@ -63,8 +61,6 @@
  nrm            :: p
  NRM            :: p
 
-local
-
 syntax "_Proof" :: "[p,o]=>prop"    ("(_ /: _)" [51, 10] 5)
 
 ML {*
--- a/src/HOL/Bali/DeclConcepts.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Bali/DeclConcepts.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -98,7 +98,7 @@
 lemma acc_modi_accmodi_simp[simp]: "accmodi (a::acc_modi) = a"
 by (simp add: acc_modi_accmodi_def)
 
-instantiation decl_ext_type:: (type) has_accmodi
+instantiation decl_ext :: (type) has_accmodi
 begin
 
 definition
@@ -150,7 +150,7 @@
 class has_declclass =
   fixes declclass:: "'a \<Rightarrow> qtname"
 
-instantiation qtname_ext_type :: (type) has_declclass
+instantiation qtname_ext :: (type) has_declclass
 begin
 
 definition
@@ -162,7 +162,7 @@
 
 lemma qtname_declclass_def:
   "declclass q \<equiv> (q::qtname)"
-  by (induct q) (simp add: declclass_qtname_ext_type_def)
+  by (induct q) (simp add: declclass_qtname_ext_def)
 
 lemma qtname_declclass_simp[simp]: "declclass (q::qtname) = q"
 by (simp add: qtname_declclass_def)
@@ -186,14 +186,14 @@
 class has_static =
   fixes is_static :: "'a \<Rightarrow> bool"
 
-instantiation decl_ext_type :: (has_static) has_static
+instantiation decl_ext :: (has_static) has_static
 begin
 
 instance ..
 
 end
 
-instantiation member_ext_type :: (type) has_static
+instantiation member_ext :: (type) has_static
 begin
 
 instance ..
@@ -457,21 +457,21 @@
 class has_resTy =
   fixes resTy:: "'a \<Rightarrow> ty"
 
-instantiation decl_ext_type :: (has_resTy) has_resTy
+instantiation decl_ext :: (has_resTy) has_resTy
 begin
 
 instance ..
 
 end
 
-instantiation member_ext_type :: (has_resTy) has_resTy
+instantiation member_ext :: (has_resTy) has_resTy
 begin
 
 instance ..
 
 end
 
-instantiation mhead_ext_type :: (type) has_resTy
+instantiation mhead_ext :: (type) has_resTy
 begin
 
 instance ..
@@ -487,7 +487,7 @@
 lemma resTy_mhead [simp]:"resTy (mhead m) = resTy m"
 by (simp add: mhead_def mhead_resTy_simp)
 
-instantiation prod :: ("type","has_resTy") has_resTy
+instantiation prod :: (type, has_resTy) has_resTy
 begin
 
 definition
--- a/src/HOL/Bali/Name.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Bali/Name.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -75,7 +75,7 @@
 end
 
 definition
-  qtname_qtname_def: "qtname (q::'a qtname_ext_type) = q"
+  qtname_qtname_def: "qtname (q::'a qtname_scheme) = q"
 
 translations
   (type) "qtname" <= (type) "\<lparr>pid::pname,tid::tname\<rparr>"
--- a/src/HOL/Bali/TypeRel.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Bali/TypeRel.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -512,12 +512,13 @@
 apply (ind_cases "G\<turnstile>S\<preceq>NT")
 by auto
 
-lemma widen_Object:"\<lbrakk>isrtype G T;ws_prog G\<rbrakk> \<Longrightarrow> G\<turnstile>RefT T \<preceq> Class Object"
-apply (case_tac T)
-apply (auto)
-apply (subgoal_tac "G\<turnstile>qtname_ext_type\<preceq>\<^sub>C Object")
-apply (auto intro: subcls_ObjectI)
-done
+lemma widen_Object:
+  assumes "isrtype G T" and "ws_prog G"
+  shows "G\<turnstile>RefT T \<preceq> Class Object"
+proof (cases T)
+  case (ClassT C) with assms have "G\<turnstile>C\<preceq>\<^sub>C Object" by (auto intro: subcls_ObjectI)
+  with ClassT show ?thesis by auto
+qed simp_all
 
 lemma widen_trans_lemma [rule_format (no_asm)]: 
   "\<lbrakk>G\<turnstile>S\<preceq>U; \<forall>C. is_class G C \<longrightarrow> G\<turnstile>C\<preceq>\<^sub>C Object\<rbrakk> \<Longrightarrow> \<forall>T. G\<turnstile>U\<preceq>T \<longrightarrow> G\<turnstile>S\<preceq>T"
--- a/src/HOL/Decision_Procs/Approximation.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -3301,11 +3301,11 @@
 ML {*
   fun reorder_bounds_tac prems i =
     let
-      fun variable_of_bound (Const ("Trueprop", _) $
+      fun variable_of_bound (Const (@{const_name Trueprop}, _) $
                              (Const (@{const_name Set.member}, _) $
                               Free (name, _) $ _)) = name
-        | variable_of_bound (Const ("Trueprop", _) $
-                             (Const ("op =", _) $
+        | variable_of_bound (Const (@{const_name Trueprop}, _) $
+                             (Const (@{const_name "op ="}, _) $
                               Free (name, _) $ _)) = name
         | variable_of_bound t = raise TERM ("variable_of_bound", [t])
 
--- a/src/HOL/Decision_Procs/Cooper.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Decision_Procs/Cooper.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -1960,12 +1960,12 @@
       @{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')
-  | fm_of_term ps vs (Const ("Ex", _) $ Abs (xn, xT, p)) =
+  | fm_of_term ps vs (Const (@{const_name Ex}, _) $ Abs (xn, xT, p)) =
       let
         val (xn', p') = variant_abs (xn, xT, p);
         val vs' = (Free (xn', xT), 0) :: map (fn (v, n) => (v, n + 1)) vs;
       in @{code E} (fm_of_term ps vs' p) end
-  | fm_of_term ps vs (Const ("All", _) $ Abs (xn, xT, p)) =
+  | fm_of_term ps vs (Const (@{const_name All}, _) $ Abs (xn, xT, p)) =
       let
         val (xn', p') = variant_abs (xn, xT, p);
         val vs' = (Free (xn', xT), 0) :: map (fn (v, n) => (v, n + 1)) vs;
--- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy	Thu Aug 19 17:41:52 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("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
+     Const(@{const_name "op ="}, _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
                             else Ferrante_Rackoff_Data.Nox
-   | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
+   | @{term "Not"}$(Const(@{const_name "op ="}, _)$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("op =",_)$_$Const(@{const_name Groups.zero},_) =>
+|  Const(@{const_name "op ="},_)$_$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("op =",_)$a$b =>
+| Const(@{const_name "op ="},_)$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("op =",_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct
+| @{term "Not"} $(Const(@{const_name "op ="},_)$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("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
+     Const(@{const_name "op ="}, _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
                             else Ferrante_Rackoff_Data.Nox
-   | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
+   | @{term "Not"}$(Const(@{const_name "op ="}, _)$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 19 17:40:44 2010 +0200
+++ b/src/HOL/Decision_Procs/Ferrack.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -2000,9 +2000,9 @@
   | 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 "Not"} $ t') = @{code NOT} (fm_of_term vs t')
-  | fm_of_term vs (Const ("Ex", _) $ Abs (xn, xT, p)) =
+  | fm_of_term vs (Const (@{const_name Ex}, _) $ Abs (xn, xT, p)) =
       @{code E} (fm_of_term (("", dummyT) :: vs) p)
-  | fm_of_term vs (Const ("All", _) $ Abs (xn, xT, p)) =
+  | fm_of_term vs (Const (@{const_name All}, _) $ Abs (xn, xT, p)) =
       @{code A} (fm_of_term (("", dummyT) ::  vs) p)
   | fm_of_term vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term @{context} t);
 
--- a/src/HOL/Decision_Procs/MIR.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Decision_Procs/MIR.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -5845,9 +5845,9 @@
       @{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 ("Ex", _) $ Abs (xn, xT, p)) =
+  | fm_of_term vs (Const (@{const_name Ex}, _) $ Abs (xn, xT, p)) =
       @{code E} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p)
-  | fm_of_term vs (Const ("All", _) $ Abs (xn, xT, p)) =
+  | fm_of_term vs (Const (@{const_name All}, _) $ Abs (xn, xT, p)) =
       @{code A} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p)
   | fm_of_term vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term @{context} t);
 
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -2960,7 +2960,7 @@
 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("op =",rrT rT);
+fun eqt rT = Const(@{const_name "op ="},rrT rT);
 fun rz rT = Const(@{const_name Groups.zero},rT);
 
 fun dest_nat t = case t of
@@ -3015,26 +3015,26 @@
 
 fun fm_of_term m m' fm = 
  case fm of
-    Const("True",_) => @{code T}
-  | Const("False",_) => @{code F}
-  | Const("Not",_)$p => @{code NOT} (fm_of_term m m' p)
-  | Const("op &",_)$p$q => @{code And} (fm_of_term m m' p, fm_of_term m m' q)
-  | Const("op |",_)$p$q => @{code Or} (fm_of_term m m' p, fm_of_term m m' q)
-  | Const("op -->",_)$p$q => @{code Imp} (fm_of_term m m' p, fm_of_term m m' q)
-  | Const("op =",ty)$p$q => 
+    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 => 
        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 => 
         @{code Lt} (@{code Sub} (tm_of_term m m' p, tm_of_term m m' q))
   | Const(@{const_name Orderings.less_eq},_)$p$q => 
         @{code Le} (@{code Sub} (tm_of_term m m' p, tm_of_term m m' q))
-  | Const("Ex",_)$Abs(xn,xT,p) => 
+  | Const(@{const_name Ex},_)$Abs(xn,xT,p) => 
      let val (xn', p') =  variant_abs (xn,xT,p)
          val x = Free(xn',xT)
          fun incr i = i + 1
          val m0 = (x,0):: (map (apsnd incr) m)
       in @{code E} (fm_of_term m0 m' p') end
-  | Const("All",_)$Abs(xn,xT,p) => 
+  | Const(@{const_name All},_)$Abs(xn,xT,p) => 
      let val (xn', p') =  variant_abs (xn,xT,p)
          val x = Free(xn',xT)
          fun incr i = i + 1
@@ -3045,8 +3045,8 @@
 
 fun term_of_fm T m m' t = 
   case t of
-    @{code T} => Const("True",bT)
-  | @{code F} => Const("False",bT)
+    @{code T} => Const(@{const_name True},bT)
+  | @{code F} => Const(@{const_name False},bT)
   | @{code NOT} p => nott $ (term_of_fm T m m' p)
   | @{code And} (p,q) => conjt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
   | @{code Or} (p,q) => disjt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
--- a/src/HOL/Decision_Procs/cooper_tac.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Decision_Procs/cooper_tac.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -110,7 +110,7 @@
     fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i)
     (* The result of the quantifier elimination *)
     val (th, tac) = case (prop_of pre_thm) of
-        Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ =>
+        Const ("==>", _) $ (Const (@{const_name Trueprop}, _) $ t1) $ _ =>
     let val pth = linzqe_oracle (cterm_of thy (Pattern.eta_long [] t1))
     in
           ((pth RS iffD2) RS pre_thm,
--- a/src/HOL/Decision_Procs/ferrack_tac.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Decision_Procs/ferrack_tac.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -91,7 +91,7 @@
     fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i)
     (* The result of the quantifier elimination *)
     val (th, tac) = case prop_of pre_thm of
-        Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ =>
+        Const ("==>", _) $ (Const (@{const_name Trueprop}, _) $ t1) $ _ =>
     let val pth = linr_oracle (ctxt, Pattern.eta_long [] t1)
     in 
           (trace_msg ("calling procedure with term:\n" ^
--- a/src/HOL/Decision_Procs/ferrante_rackoff.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Decision_Procs/ferrante_rackoff.ML	Thu Aug 19 17:41:52 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("op &", _)$ _ $ _ =>
+   Const(@{const_name "op &"}, _)$ _ $ _ =>
      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("op |", _)$ _ $ _ =>
+ |  Const(@{const_name "op |"}, _)$ _ $ _ =>
      let
        val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb
        val (lS,lth) = uset vars l  val (rS, rth) = uset vars r
@@ -77,7 +77,7 @@
  fun main vs p =
   let
    val ((xn,ce),(x,fm)) = (case term_of p of
-                   Const("Ex",_)$Abs(xn,xT,_) =>
+                   Const(@{const_name Ex},_)$Abs(xn,xT,_) =>
                         Thm.dest_comb p ||> Thm.dest_abs (SOME xn) |>> pair xn
                  | _ => raise CTERM ("main QE only treats existential quantifiers!", [p]))
    val cT = ctyp_of_term x
@@ -122,12 +122,12 @@
 
    fun decomp_mpinf fm =
      case term_of fm of
-       Const("op &",_)$_$_ =>
+       Const(@{const_name "op &"},_)$_$_ =>
         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("op |",_)$_$_ =>
+     | Const(@{const_name "op |"},_)$_$_ =>
         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,19 +175,19 @@
  let
   fun h bounds tm =
    (case term_of tm of
-     Const ("op =", T) $ _ $ _ =>
+     Const (@{const_name "op ="}, T) $ _ $ _ =>
        if domain_type T = HOLogic.boolT then find_args bounds tm
        else Thm.dest_fun2 tm
-   | Const ("Not", _) $ _ => h bounds (Thm.dest_arg tm)
-   | Const ("All", _) $ _ => find_body bounds (Thm.dest_arg tm)
-   | Const ("Ex", _) $ _ => find_body bounds (Thm.dest_arg tm)
-   | Const ("op &", _) $ _ $ _ => find_args bounds tm
-   | Const ("op |", _) $ _ $ _ => find_args bounds tm
-   | Const ("op -->", _) $ _ $ _ => find_args bounds 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 ("==>", _) $ _ $ _ => find_args bounds tm
    | Const ("==", _) $ _ $ _ => find_args bounds tm
    | Const ("all", _) $ _ => find_body bounds (Thm.dest_arg tm)
-   | Const ("Trueprop", _) $ _ => h bounds (Thm.dest_arg tm)
+   | Const (@{const_name Trueprop}, _) $ _ => h bounds (Thm.dest_arg tm)
    | _ => Thm.dest_fun2 tm)
   and find_args bounds tm =
            (h bounds (Thm.dest_arg tm) handle CTERM _ => Thm.dest_arg1 tm)
--- a/src/HOL/Decision_Procs/langford.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Decision_Procs/langford.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -30,7 +30,7 @@
 
 fun basic_dloqe ctxt stupid dlo_qeth dlo_qeth_nolb dlo_qeth_noub gather ep = 
  case term_of ep of 
-  Const("Ex",_)$_ => 
+  Const(@{const_name Ex},_)$_ => 
    let 
      val p = Thm.dest_arg ep
      val ths = simplify (HOL_basic_ss addsimps gather) (instantiate' [] [SOME p] stupid)
@@ -116,13 +116,13 @@
 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("op =",_)$l$r => l aconv term_of x orelse r aconv term_of x
+   Const(@{const_name "op ="},_)$l$r => l aconv term_of x orelse r aconv term_of x
  | _ => false ;
 
 local 
 fun proc ct = 
  case term_of ct of
-  Const("Ex",_)$Abs (xn,_,_) => 
+  Const(@{const_name Ex},_)$Abs (xn,_,_) => 
    let 
     val e = Thm.dest_fun ct
     val (x,p) = Thm.dest_abs (SOME xn) (Thm.dest_arg ct)
@@ -176,19 +176,19 @@
  let
   fun h bounds tm =
    (case term_of tm of
-     Const ("op =", T) $ _ $ _ =>
+     Const (@{const_name "op ="}, T) $ _ $ _ =>
        if domain_type T = HOLogic.boolT then find_args bounds tm
        else Thm.dest_fun2 tm
-   | Const ("Not", _) $ _ => h bounds (Thm.dest_arg tm)
-   | Const ("All", _) $ _ => find_body bounds (Thm.dest_arg 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 ("Ex", _) $ _ => find_body bounds (Thm.dest_arg tm)
-   | Const ("op &", _) $ _ $ _ => find_args bounds tm
-   | Const ("op |", _) $ _ $ _ => find_args bounds tm
-   | Const ("op -->", _) $ _ $ _ => find_args bounds 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 ("==>", _) $ _ $ _ => find_args bounds tm
    | Const ("==", _) $ _ $ _ => find_args bounds tm
-   | Const ("Trueprop", _) $ _ => h bounds (Thm.dest_arg tm)
+   | Const (@{const_name Trueprop}, _) $ _ => h bounds (Thm.dest_arg tm)
    | _ => Thm.dest_fun2 tm)
   and find_args bounds tm =
     (h bounds (Thm.dest_arg tm) handle CTERM _ => h bounds (Thm.dest_arg1 tm))
--- a/src/HOL/Decision_Procs/mir_tac.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Decision_Procs/mir_tac.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -132,7 +132,7 @@
     fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i)
     (* The result of the quantifier elimination *)
     val (th, tac) = case (prop_of pre_thm) of
-        Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ =>
+        Const ("==>", _) $ (Const (@{const_name Trueprop}, _) $ t1) $ _ =>
     let val pth =
           (* If quick_and_dirty then run without proof generation as oracle*)
              if !quick_and_dirty
--- a/src/HOL/HOL.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/HOL.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -47,33 +47,28 @@
   "fun" :: (type, type) type
   itself :: (type) type
 
-global
-
 typedecl bool
 
 judgment
   Trueprop      :: "bool => prop"                   ("(_)" 5)
 
 consts
-  Not           :: "bool => bool"                   ("~ _" [40] 40)
   True          :: bool
   False         :: bool
+  Not           :: "bool => bool"                   ("~ _" [40] 40)
 
+global consts
+  "op &"        :: "[bool, bool] => bool"           (infixr "&" 35)
+  "op |"        :: "[bool, bool] => bool"           (infixr "|" 30)
+  "op -->"      :: "[bool, bool] => bool"           (infixr "-->" 25)
+
+  "op ="        :: "['a, 'a] => bool"               (infixl "=" 50)
+
+local consts
   The           :: "('a => bool) => 'a"
   All           :: "('a => bool) => bool"           (binder "ALL " 10)
   Ex            :: "('a => bool) => bool"           (binder "EX " 10)
   Ex1           :: "('a => bool) => bool"           (binder "EX! " 10)
-  Let           :: "['a, 'a => 'b] => 'b"
-
-  "op ="        :: "['a, 'a] => bool"               (infixl "=" 50)
-  "op &"        :: "[bool, bool] => bool"           (infixr "&" 35)
-  "op |"        :: "[bool, bool] => bool"           (infixr "|" 30)
-  "op -->"      :: "[bool, bool] => bool"           (infixr "-->" 25)
-
-local
-
-consts
-  If            :: "[bool, 'a, 'a] => 'a"           ("(if (_)/ then (_)/ else (_))" [0, 0, 10] 10)
 
 
 subsubsection {* Additional concrete syntax *}
@@ -127,8 +122,6 @@
 
 translations
   "THE x. P"              == "CONST The (%x. P)"
-  "_Let (_binds b bs) e"  == "_Let b (_Let bs e)"
-  "let x = a in e"        == "CONST Let a (%x. e)"
 
 print_translation {*
   [(@{const_syntax The}, fn [Abs abs] =>
@@ -185,15 +178,21 @@
   iff:          "(P-->Q) --> (Q-->P) --> (P=Q)"
   True_or_False:  "(P=True) | (P=False)"
 
-defs
-  Let_def [code]: "Let s f == f(s)"
-  if_def:         "If P x y == THE z::'a. (P=True --> z=x) & (P=False --> z=y)"
-
 finalconsts
   "op ="
   "op -->"
   The
 
+definition If :: "bool \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("(if (_)/ then (_)/ else (_))" [0, 0, 10] 10) where
+  "If P x y \<equiv> (THE z::'a. (P=True --> z=x) & (P=False --> z=y))"
+
+definition Let :: "'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b" where
+  "Let s f \<equiv> f s"
+
+translations
+  "_Let (_binds b bs) e"  == "_Let b (_Let bs e)"
+  "let x = a in e"        == "CONST Let a (%x. e)"
+
 axiomatization
   undefined :: 'a
 
@@ -1084,16 +1083,16 @@
 text {* \medskip if-then-else rules *}
 
 lemma if_True [code]: "(if True then x else y) = x"
-  by (unfold if_def) blast
+  by (unfold If_def) blast
 
 lemma if_False [code]: "(if False then x else y) = y"
-  by (unfold if_def) blast
+  by (unfold If_def) blast
 
 lemma if_P: "P ==> (if P then x else y) = x"
-  by (unfold if_def) blast
+  by (unfold If_def) blast
 
 lemma if_not_P: "~P ==> (if P then x else y) = y"
-  by (unfold if_def) blast
+  by (unfold If_def) blast
 
 lemma split_if: "P (if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))"
   apply (rule case_split [of Q])
@@ -1365,7 +1364,7 @@
       and "c \<Longrightarrow> x = u"
       and "\<not> c \<Longrightarrow> y = v"
   shows "(if b then x else y) = (if c then u else v)"
-  unfolding if_def using assms by simp
+  using assms by simp
 
 text {* Prevents simplification of x and y:
   faster and allows the execution of functional programs. *}
@@ -1390,13 +1389,6 @@
   "f (if c then x else y) = (if c then f x else f y)"
   by simp
 
-text {* This lemma restricts the effect of the rewrite rule u=v to the left-hand
-  side of an equality.  Used in @{text "{Integ,Real}/simproc.ML"} *}
-lemma restrict_to_left:
-  assumes "x = y"
-  shows "(x = z) = (y = z)"
-  using assms by simp
-
 
 subsubsection {* Generic cases and induction *}
 
@@ -1582,7 +1574,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 "op |"}, @{const_name "op &"}, @{const_name Ex}]
 );
 *}
 
--- a/src/HOL/Import/HOL/bool.imp	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Import/HOL/bool.imp	Thu Aug 19 17:41:52 2010 +0200
@@ -12,11 +12,11 @@
   "ARB" > "ARB_def"
 
 const_maps
-  "~" > "Not"
+  "~" > "HOL.Not"
   "bool_case" > "Datatype.bool.bool_case"
   "\\/" > "op |"
   "TYPE_DEFINITION" > "HOL4Setup.TYPE_DEFINITION"
-  "T" > "True"
+  "T" > "HOL.True"
   "RES_SELECT" > "HOL4Base.bool.RES_SELECT"
   "RES_FORALL" > "HOL4Base.bool.RES_FORALL"
   "RES_EXISTS_UNIQUE" > "HOL4Base.bool.RES_EXISTS_UNIQUE"
@@ -25,13 +25,13 @@
   "ONE_ONE" > "HOL4Setup.ONE_ONE"
   "LET" > "HOL4Compat.LET"
   "IN" > "HOL4Base.bool.IN"
-  "F" > "False"
+  "F" > "HOL.False"
   "COND" > "HOL.If"
   "ARB" > "HOL4Base.bool.ARB"
-  "?!" > "Ex1"
-  "?" > "Ex"
+  "?!" > "HOL.Ex1"
+  "?" > "HOL.Ex"
   "/\\" > "op &"
-  "!" > "All"
+  "!" > "HOL.All"
 
 thm_maps
   "bool_case_thm" > "HOL4Base.bool.bool_case_thm"
--- a/src/HOL/Import/HOLLight/hollight.imp	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Import/HOLLight/hollight.imp	Thu Aug 19 17:41:52 2010 +0200
@@ -18,13 +18,13 @@
   "finite_sum" > "HOLLight.hollight.finite_sum"
   "finite_image" > "HOLLight.hollight.finite_image"
   "cart" > "HOLLight.hollight.cart"
-  "bool" > "bool"
+  "bool" > "HOL.bool"
   "N_3" > "HOLLight.hollight.N_3"
   "N_2" > "HOLLight.hollight.N_2"
   "N_1" > "Product_Type.unit"
 
 const_maps
-  "~" > "Not"
+  "~" > "HOL.Not"
   "two_2" > "HOLLight.hollight.two_2"
   "two_1" > "HOLLight.hollight.two_1"
   "treal_of_num" > "HOLLight.hollight.treal_of_num"
@@ -229,8 +229,8 @@
   "ALL2" > "HOLLight.hollight.ALL2"
   "ABS_prod" > "Abs_Prod"
   "@" > "Hilbert_Choice.Eps"
-  "?!" > "Ex1"
-  "?" > "Ex"
+  "?!" > "HOL.Ex1"
+  "?" > "HOL.Ex"
   ">=" > "HOLLight.hollight.>="
   ">" > "HOLLight.hollight.>"
   "==>" > "op -->"
@@ -243,7 +243,7 @@
   "+" > "Groups.plus" :: "nat => nat => nat"
   "*" > "Groups.times" :: "nat => nat => nat"
   "$" > "HOLLight.hollight.$"
-  "!" > "All"
+  "!" > "HOL.All"
 
 const_renames
   "ALL" > "ALL_list"
--- a/src/HOL/Import/hol4rews.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Import/hol4rews.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -616,7 +616,7 @@
     fun handle_meta [x as Appl[Appl[Constant "_constrain", Constant @{const_syntax "=="}, _],_,_]] = x
       | handle_meta [x as Appl[Appl[Constant "_constrain", Constant @{const_syntax all}, _],_]] = x
       | handle_meta [x as Appl[Appl[Constant "_constrain", Constant @{const_syntax "==>"}, _],_,_]] = x
-      | handle_meta [x] = Appl[Constant "Trueprop",x]
+      | handle_meta [x] = Appl[Constant @{const_syntax Trueprop}, x]
       | handle_meta _ = error "hol4rews error: Trueprop not applied to single argument"
 in
 val smarter_trueprop_parsing = [(@{const_syntax Trueprop},handle_meta)]
@@ -624,12 +624,12 @@
 
 local
     fun initial_maps thy =
-        thy |> add_hol4_type_mapping "min" "bool" false "bool"
+        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 "Nat.ind"
-            |> add_hol4_const_mapping "min" "=" false "op ="
-            |> add_hol4_const_mapping "min" "==>" false "op -->"
-            |> add_hol4_const_mapping "min" "@" false "Hilbert_Choice.Eps"
+            |> 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 "Eps"}
 in
 val hol4_setup =
   initial_maps #>
--- a/src/HOL/Import/proof_kernel.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Import/proof_kernel.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -388,7 +388,6 @@
 
 val type_of = Term.type_of
 
-val boolT = Type("bool",[])
 val propT = Type("prop",[])
 
 fun mk_defeq name rhs thy =
@@ -1007,7 +1006,7 @@
 local
     val th = thm "not_def"
     val thy = theory_of_thm th
-    val pp = Thm.reflexive (cterm_of thy (Const("Trueprop",boolT-->propT)))
+    val pp = Thm.reflexive (cterm_of thy (Const(@{const_name Trueprop},HOLogic.boolT-->propT)))
 in
 val not_elim_thm = Thm.combination pp th
 end
@@ -1053,9 +1052,9 @@
         val c = prop_of th3
         val vname = fst(dest_Free v)
         val (cold,cnew) = case c of
-                              tpc $ (Const("All",allT) $ Abs(oldname,ty,body)) =>
+                              tpc $ (Const(@{const_name All},allT) $ Abs(oldname,ty,body)) =>
                               (Abs(oldname,dummyT,Bound 0),Abs(vname,dummyT,Bound 0))
-                            | tpc $ (Const("All",allT) $ rest) => (tpc,tpc)
+                            | tpc $ (Const(@{const_name All},allT) $ rest) => (tpc,tpc)
                             | _ => raise ERR "mk_GEN" "Unknown conclusion"
         val th4 = Thm.rename_boundvars cold cnew th3
         val res = implies_intr_hyps th4
@@ -1205,7 +1204,8 @@
 
 fun non_trivial_term_consts t = fold_aterms
   (fn Const (c, _) =>
-      if c = "Trueprop" orelse c = "All" orelse c = "op -->" orelse c = "op &" orelse c = "op ="
+      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 ="}
       then I else insert (op =) c
     | _ => I) t [];
 
@@ -1213,12 +1213,12 @@
     let
         fun add_consts (Const (c, _), cs) =
             (case c of
-                 "op =" => insert (op =) "==" cs
-               | "op -->" => insert (op =) "==>" cs
-               | "All" => cs
+                 @{const_name "op ="} => insert (op =) "==" cs
+               | @{const_name "op -->"} => insert (op =) "==>" cs
+               | @{const_name All} => cs
                | "all" => cs
-               | "op &" => cs
-               | "Trueprop" => cs
+               | @{const_name "op &"} => cs
+               | @{const_name Trueprop} => cs
                | _ => insert (op =) c cs)
           | add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs))
           | add_consts (Abs (_, _, t), cs) = add_consts (t, cs)
@@ -1476,10 +1476,10 @@
 fun mk_COMB th1 th2 thy =
     let
         val (f,g) = case concl_of th1 of
-                        _ $ (Const("op =",_) $ f $ g) => (f,g)
+                        _ $ (Const(@{const_name "op ="},_) $ f $ g) => (f,g)
                       | _ => raise ERR "mk_COMB" "First theorem not an equality"
         val (x,y) = case concl_of th2 of
-                        _ $ (Const("op =",_) $ x $ y) => (x,y)
+                        _ $ (Const(@{const_name "op ="},_) $ 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("op |",_) $ l $ r) => (l,r)
+                        _ $ (Const(@{const_name "op |"},_) $ 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
@@ -1654,7 +1654,7 @@
         val cwit = cterm_of thy wit'
         val cty = ctyp_of_term cwit
         val a = case ex' of
-                    (Const("Ex",_) $ a) => a
+                    (Const(@{const_name Ex},_) $ a) => a
                   | _ => raise ERR "EXISTS" "Argument not existential"
         val ca = cterm_of thy a
         val exists_thm' = beta_eta_thm (Drule.instantiate' [SOME cty] [SOME ca,SOME cwit] exists_thm)
@@ -1687,7 +1687,7 @@
         val c = HOLogic.dest_Trueprop (concl_of th2)
         val cc = cterm_of thy c
         val a = case concl_of th1 of
-                    _ $ (Const("Ex",_) $ a) => a
+                    _ $ (Const(@{const_name Ex},_) $ a) => a
                   | _ => raise ERR "CHOOSE" "Conclusion not existential"
         val ca = cterm_of (theory_of_thm th1) a
         val choose_thm' = beta_eta_thm (Drule.instantiate' [SOME cvty] [SOME ca,SOME cc] choose_thm)
@@ -1773,7 +1773,7 @@
         val (info',tm') = disamb_term_from info tm
         val th = norm_hyps th
         val ct = cterm_of thy tm'
-        val th1 = rearrange thy (HOLogic.mk_Trueprop (Const("Not",boolT-->boolT) $ tm')) th
+        val th1 = rearrange thy (HOLogic.mk_Trueprop (Const(@{const_name Not},HOLogic.boolT-->HOLogic.boolT) $ tm')) th
         val ccontr_thm' = Drule.instantiate' [] [SOME ct] ccontr_thm
         val res1 = uniq_compose ((nprems_of th1) - 1) th1 1 ccontr_thm'
         val res = HOLThm(rens_of info',res1)
@@ -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("op =",_) $ f $ g) => (Term.lambda v f,Term.lambda v g)
+                        _ $ (Const(@{const_name "op ="},_) $ 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("op -->",_) $ a $ Const("False",_)) => a
+                    _ $ (Const(@{const_name "op -->"},_) $ 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
@@ -1877,7 +1877,7 @@
         val _ = if_debug pth hth
         val th1 = implies_elim_all (beta_eta_thm th)
         val a = case concl_of th1 of
-                    _ $ (Const("Not",_) $ a) => a
+                    _ $ (Const(@{const_name Not},_) $ a) => a
                   | _ => raise ERR "NOT_ELIM" "Conclusion of bad form"
         val ca = cterm_of thy a
         val th2 = Thm.equal_elim (Drule.instantiate' [] [SOME ca] not_elim_thm) th1
@@ -1996,7 +1996,7 @@
                                        ("x",dT,body $ Bound 0)
                                    end
                                    handle TYPE _ => raise ERR "new_specification" "not an abstraction type"
-                               fun dest_exists (Const("Ex",_) $ abody) =
+                               fun dest_exists (Const(@{const_name Ex},_) $ abody) =
                                    dest_eta_abs abody
                                  | dest_exists tm =
                                    raise ERR "new_specification" "Bad existential formula"
@@ -2082,7 +2082,7 @@
             val (HOLThm(rens,td_th)) = norm_hthm thy hth
             val th2 = beta_eta_thm (td_th RS ex_imp_nonempty)
             val c = case concl_of th2 of
-                        _ $ (Const("Ex",_) $ Abs(_,_,Const(@{const_name Set.member},_) $ _ $ c)) => c
+                        _ $ (Const(@{const_name Ex},_) $ Abs(_,_,Const(@{const_name Set.member},_) $ _ $ c)) => c
                       | _ => raise ERR "new_type_definition" "Bad type definition theorem"
             val tfrees = OldTerm.term_tfrees c
             val tnames = map fst tfrees
@@ -2108,7 +2108,7 @@
             val rew = rewrite_hol4_term (concl_of td_th) thy4
             val th  = Thm.equal_elim rew (Thm.transfer thy4 td_th)
             val c   = case HOLogic.dest_Trueprop (prop_of th) of
-                          Const("Ex",exT) $ P =>
+                          Const(@{const_name Ex},exT) $ P =>
                           let
                               val PT = domain_type exT
                           in
@@ -2157,7 +2157,7 @@
                                     SOME (cterm_of thy t)] light_nonempty
             val th2 = beta_eta_thm (td_th RS (beta_eta_thm light_nonempty'))
             val c = case concl_of th2 of
-                        _ $ (Const("Ex",_) $ Abs(_,_,Const(@{const_name Set.member},_) $ _ $ c)) => c
+                        _ $ (Const(@{const_name Ex},_) $ Abs(_,_,Const(@{const_name Set.member},_) $ _ $ c)) => c
                       | _ => raise ERR "type_introduction" "Bad type definition theorem"
             val tfrees = OldTerm.term_tfrees c
             val tnames = sort_strings (map fst tfrees)
--- a/src/HOL/Import/shuffler.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Import/shuffler.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -60,14 +60,14 @@
 
 fun mk_meta_eq th =
     (case concl_of th of
-         Const("Trueprop",_) $ (Const("op =",_) $ _ $ _) => th RS eq_reflection
+         Const(@{const_name Trueprop},_) $ (Const(@{const_name "op ="},_) $ _ $ _) => 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("Trueprop",_) $ (Const("op =",_) $ _ $ _) => th
+         Const(@{const_name Trueprop},_) $ (Const(@{const_name "op ="},_) $ _ $ _) => 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/IsaMakefile	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/IsaMakefile	Thu Aug 19 17:41:52 2010 +0200
@@ -213,7 +213,6 @@
   Tools/sat_funcs.ML \
   Tools/sat_solver.ML \
   Tools/split_rule.ML \
-  Tools/typedef_codegen.ML \
   Tools/typedef.ML \
   Transitive_Closure.thy \
   Typedef.thy \
@@ -303,7 +302,6 @@
   Tools/Predicate_Compile/predicate_compile_specialisation.ML \
   Tools/Predicate_Compile/predicate_compile_pred.ML \
   Tools/quickcheck_generators.ML \
-  Tools/quickcheck_record.ML \
   Tools/Qelim/cooper.ML \
   Tools/Qelim/cooper_procedure.ML \
   Tools/Qelim/qelim.ML \
@@ -343,7 +341,6 @@
   Tools/string_code.ML \
   Tools/string_syntax.ML \
   Tools/transfer.ML \
-  Tools/typecopy.ML \
   Tools/TFL/casesplit.ML \
   Tools/TFL/dcterm.ML \
   Tools/TFL/post.ML \
--- a/src/HOL/Library/Dlist.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Library/Dlist.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -15,8 +15,8 @@
 qed
 
 lemma dlist_ext:
-  assumes "list_of_dlist xs = list_of_dlist ys"
-  shows "xs = ys"
+  assumes "list_of_dlist dxs = list_of_dlist dys"
+  shows "dxs = dys"
   using assms by (simp add: list_of_dlist_inject)
 
 
@@ -107,6 +107,19 @@
   by simp
 
 
+text {* Equality *}
+
+instantiation dlist :: (eq) eq
+begin
+
+definition "HOL.eq dxs dys \<longleftrightarrow> HOL.eq (list_of_dlist dxs) (list_of_dlist dys)"
+
+instance proof
+qed (simp add: eq_dlist_def eq list_of_dlist_inject)
+
+end
+
+
 section {* Induction principle and case distinction *}
 
 lemma dlist_induct [case_names empty insert, induct type: dlist]:
@@ -283,6 +296,7 @@
 
 end
 
+
 hide_const (open) member fold foldr empty insert remove map filter null member length fold
 
 end
--- a/src/HOL/Library/Eval_Witness.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Library/Eval_Witness.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -63,7 +63,7 @@
     else error ("Type " ^ quote (Syntax.string_of_typ_global thy T) ^ " not allowed for ML witnesses")
 
   fun dest_exs  0 t = t
-    | dest_exs n (Const ("Ex", _) $ Abs (v,T,b)) = 
+    | dest_exs n (Const (@{const_name Ex}, _) $ Abs (v,T,b)) = 
       Abs (v, check_type T, dest_exs (n - 1) b)
     | dest_exs _ _ = sys_error "dest_exs";
   val t = dest_exs (length ws) (HOLogic.dest_Trueprop goal);
--- a/src/HOL/Library/ROOT.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Library/ROOT.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -2,4 +2,4 @@
 (* Classical Higher-order Logic -- batteries included *)
 
 use_thys ["Library", "List_Prefix", "List_lexord", "Sublist_Order",
-  "Code_Char_chr", "Code_Integer", "Efficient_Nat", "Executable_Set", "Code_Prolog"];
+  "Code_Char_chr", "Code_Integer", "Efficient_Nat", "Executable_Set"(*, "Code_Prolog"*)];
--- a/src/HOL/Library/positivstellensatz.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Library/positivstellensatz.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -498,7 +498,7 @@
  val strip_exists =
   let fun h (acc, t) =
    case (term_of t) of
-    Const("Ex",_)$Abs(x,T,p) => h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc))
+    Const(@{const_name Ex},_)$Abs(x,T,p) => h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc))
   | _ => (acc,t)
   in fn t => h ([],t)
   end
@@ -515,7 +515,7 @@
  fun mk_ex v t = Thm.capply (ext (ctyp_of_term v)) (Thm.cabs v t)
 
  fun choose v th th' = case concl_of th of 
-   @{term Trueprop} $ (Const("Ex",_)$_) => 
+   @{term Trueprop} $ (Const(@{const_name Ex},_)$_) => 
     let
      val p = (funpow 2 Thm.dest_arg o cprop_of) th
      val T = (hd o Thm.dest_ctyp o ctyp_of_term) p
@@ -533,7 +533,7 @@
  val strip_forall =
   let fun h (acc, t) =
    case (term_of t) of
-    Const("All",_)$Abs(x,T,p) => h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc))
+    Const(@{const_name All},_)$Abs(x,T,p) => h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc))
   | _ => (acc,t)
   in fn t => h ([],t)
   end
--- a/src/HOL/Library/reflection.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Library/reflection.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -82,7 +82,7 @@
 fun rearrange congs =
   let
     fun P (_, th) =
-      let val @{term "Trueprop"}$(Const ("op =",_) $l$_) = concl_of th
+      let val @{term "Trueprop"}$(Const (@{const_name "op ="},_) $l$_) = concl_of th
       in can dest_Var l end
     val (yes,no) = List.partition P congs
   in no @ yes end
--- a/src/HOL/Matrix/Matrix.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Matrix/Matrix.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -528,12 +528,12 @@
       show "f (s 0) (foldseq f (\<lambda>k. if k < m then s (Suc k) else foldseq f (\<lambda>k. s (Suc (k + m))) (na - m)) m) =
         foldseq f (\<lambda>k. if k < m then s k else foldseq f (\<lambda>k. s (k + m)) (Suc na - m)) m"
         apply (simp add: subd)
-        apply (case_tac "m=0")
+        apply (cases "m = 0")
         apply (simp)
         apply (drule sube)
         apply (auto)
         apply (rule a)
-        by (simp add: subc if_def)
+        by (simp add: subc cong del: if_cong)
     qed
   then show ?thesis using assms by simp
 qed
--- a/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -174,9 +174,9 @@
   Xcoord :: int
   Ycoord :: int
 
-lemma "make_point_ext_type (dest_point_ext_type a) = a"
+lemma "Abs_point_ext (Rep_point_ext a) = a"
 nitpick [expect = none]
-by (rule dest_point_ext_type_inverse)
+by (fact Rep_point_ext_inverse)
 
 lemma "Fract a b = of_int a / of_int b"
 nitpick [card = 1, expect = none]
--- a/src/HOL/Nominal/nominal_datatype.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Nominal/nominal_datatype.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -183,7 +183,7 @@
   end;
 
 fun mk_not_sym ths = maps (fn th => case prop_of th of
-    _ $ (Const ("Not", _) $ (Const ("op =", _) $ _ $ _)) => [th, th RS not_sym]
+    _ $ (Const (@{const_name Not}, _) $ (Const (@{const_name "op ="}, _) $ _ $ _)) => [th, th RS not_sym]
   | _ => [th]) ths;
 
 fun fresh_const T U = Const ("Nominal.fresh", T --> U --> HOLogic.boolT);
@@ -1372,7 +1372,7 @@
                             SUBGOAL (fn (t, i) => case Logic.strip_assums_concl t of
                                 _ $ (Const ("Nominal.fresh", _) $ _ $ _) =>
                                   simp_tac ind_ss1' i
-                              | _ $ (Const ("Not", _) $ _) =>
+                              | _ $ (Const (@{const_name Not}, _) $ _) =>
                                   resolve_tac freshs2' i
                               | _ => asm_simp_tac (HOL_basic_ss addsimps
                                   pt2_atoms addsimprocs [perm_simproc]) i)) 1])
@@ -1671,7 +1671,7 @@
     val rec_unique_frees' =
       Datatype_Prop.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts;
     val rec_unique_concls = map (fn ((x, U), R) =>
-        Const ("Ex1", (U --> HOLogic.boolT) --> HOLogic.boolT) $
+        Const (@{const_name Ex1}, (U --> HOLogic.boolT) --> HOLogic.boolT) $
           Abs ("y", U, R $ Free x $ Bound 0))
       (rec_unique_frees ~~ rec_result_Ts ~~ rec_sets);
 
@@ -1777,7 +1777,7 @@
                       | _ => false) prems';
                     val fresh_prems = filter (fn th => case prop_of th of
                         _ $ (Const ("Nominal.fresh", _) $ _ $ _) => true
-                      | _ $ (Const ("Not", _) $ _) => true
+                      | _ $ (Const (@{const_name Not}, _) $ _) => true
                       | _ => false) prems';
                     val Ts = map fastype_of boundsl;
 
@@ -1879,7 +1879,7 @@
                       end) rec_prems2;
 
                     val ihs = filter (fn th => case prop_of th of
-                      _ $ (Const ("All", _) $ _) => true | _ => false) prems';
+                      _ $ (Const (@{const_name All}, _) $ _) => true | _ => false) prems';
 
                     (** pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs **)
                     val _ = warning "step 7: pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs";
@@ -2022,7 +2022,7 @@
           (reccomb_names ~~ recTs ~~ rec_result_Ts))
       |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
           (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T,
-           Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
+           Const (@{const_name The}, (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
              set $ Free ("x", T) $ Free ("y", T'))))))
                (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts));
 
--- a/src/HOL/Nominal/nominal_inductive.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Nominal/nominal_inductive.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -71,14 +71,14 @@
   | 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 ("op &", _) $ p $ q) _ = (case head_of p of
+fun split_conj f names (Const (@{const_name "op &"}, _) $ p $ q) _ = (case head_of p of
       Const (name, _) =>
         if member (op =) names name then SOME (f p q) else NONE
     | _ => NONE)
   | split_conj _ _ _ _ = NONE;
 
 fun strip_all [] t = t
-  | strip_all (_ :: xs) (Const ("All", _) $ Abs (s, T, t)) = strip_all xs t;
+  | strip_all (_ :: xs) (Const (@{const_name All}, _) $ Abs (s, T, t)) = strip_all xs t;
 
 (*********************************************************************)
 (* maps  R ... & (ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t))  *)
@@ -89,7 +89,7 @@
 (* where "id" protects the subformula from simplification            *)
 (*********************************************************************)
 
-fun inst_conj_all names ps pis (Const ("op &", _) $ p $ q) _ =
+fun inst_conj_all names ps pis (Const (@{const_name "op &"}, _) $ 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 19 17:40:44 2010 +0200
+++ b/src/HOL/Nominal/nominal_inductive2.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -75,14 +75,14 @@
   | 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 ("op &", _) $ p $ q) _ = (case head_of p of
+fun split_conj f names (Const (@{const_name "op &"}, _) $ p $ q) _ = (case head_of p of
       Const (name, _) =>
         if member (op =) names name then SOME (f p q) else NONE
     | _ => NONE)
   | split_conj _ _ _ _ = NONE;
 
 fun strip_all [] t = t
-  | strip_all (_ :: xs) (Const ("All", _) $ Abs (s, T, t)) = strip_all xs t;
+  | strip_all (_ :: xs) (Const (@{const_name All}, _) $ Abs (s, T, t)) = strip_all xs t;
 
 (*********************************************************************)
 (* maps  R ... & (ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t))  *)
@@ -93,7 +93,7 @@
 (* where "id" protects the subformula from simplification            *)
 (*********************************************************************)
 
-fun inst_conj_all names ps pis (Const ("op &", _) $ p $ q) _ =
+fun inst_conj_all names ps pis (Const (@{const_name "op &"}, _) $ 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 19 17:40:44 2010 +0200
+++ b/src/HOL/Nominal/nominal_thmdecls.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -135,7 +135,7 @@
     val thy = Context.theory_of context
     val thms_to_be_added = (case (prop_of orig_thm) of
         (* case: eqvt-lemma is of the implicational form *)
-        (Const("==>", _) $ (Const ("Trueprop",_) $ hyp) $ (Const ("Trueprop",_) $ concl)) =>
+        (Const("==>", _) $ (Const (@{const_name Trueprop},_) $ hyp) $ (Const (@{const_name Trueprop},_) $ concl)) =>
           let
             val (pi,typi) = get_pi concl thy
           in
@@ -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 "op ="}, _) $
             (Const (@{const_name "perm"},typrm) $ Var (pi,typi) $ lhs) $ rhs)) =>
            (if (apply_pi lhs (pi,typi)) = rhs
                then [orig_thm]
--- a/src/HOL/Prolog/prolog.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Prolog/prolog.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -10,17 +10,17 @@
 exception not_HOHH;
 
 fun isD t = case t of
-    Const("Trueprop",_)$t     => isD t
-  | Const("op &"  ,_)$l$r     => isD l andalso isD r
-  | Const("op -->",_)$l$r     => isG l andalso isD r
+    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(   "==>",_)$l$r     => isG l andalso isD r
-  | Const("All",_)$Abs(s,_,t) => isD t
+  | Const(@{const_name All},_)$Abs(s,_,t) => isD t
   | Const("all",_)$Abs(s,_,t) => isD t
-  | Const("op |",_)$_$_       => false
-  | Const("Ex" ,_)$_          => false
-  | Const("not",_)$_          => false
-  | Const("True",_)           => false
-  | Const("False",_)          => false
+  | Const(@{const_name "op |"},_)$_$_       => false
+  | Const(@{const_name Ex} ,_)$_          => false
+  | Const(@{const_name Not},_)$_          => false
+  | Const(@{const_name True},_)           => false
+  | Const(@{const_name False},_)          => false
   | l $ r                     => isD l
   | Const _ (* rigid atom *)  => true
   | Bound _ (* rigid atom *)  => true
@@ -29,17 +29,17 @@
             anything else *)  => false
 and
     isG t = case t of
-    Const("Trueprop",_)$t     => isG t
-  | Const("op &"  ,_)$l$r     => isG l andalso isG r
-  | Const("op |"  ,_)$l$r     => isG l andalso isG r
-  | Const("op -->",_)$l$r     => isD l andalso isG r
+    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(   "==>",_)$l$r     => isD l andalso isG r
-  | Const("All",_)$Abs(_,_,t) => isG t
+  | Const(@{const_name All},_)$Abs(_,_,t) => isG t
   | Const("all",_)$Abs(_,_,t) => isG t
-  | Const("Ex" ,_)$Abs(_,_,t) => isG t
-  | Const("True",_)           => true
-  | Const("not",_)$_          => false
-  | Const("False",_)          => false
+  | Const(@{const_name Ex} ,_)$Abs(_,_,t) => isG t
+  | Const(@{const_name True},_)           => true
+  | Const(@{const_name Not},_)$_          => false
+  | Const(@{const_name False},_)          => false
   | _ (* atom *)              => true;
 
 val check_HOHH_tac1 = PRIMITIVE (fn thm =>
@@ -51,10 +51,10 @@
 
 fun atomizeD ctxt thm = let
     fun at  thm = case concl_of thm of
-      _$(Const("All" ,_)$Abs(s,_,_))=> at(thm RS
+      _$(Const(@{const_name All} ,_)$Abs(s,_,_))=> at(thm RS
         (read_instantiate ctxt [(("x", 0), "?" ^ (if s="P" then "PP" else s))] spec))
-    | _$(Const("op &",_)$_$_)       => at(thm RS conjunct1)@at(thm RS conjunct2)
-    | _$(Const("op -->",_)$_$_)     => at(thm RS mp)
+    | _$(Const(@{const_name "op &"},_)$_$_)       => at(thm RS conjunct1)@at(thm RS conjunct2)
+    | _$(Const(@{const_name "op -->"},_)$_$_)     => at(thm RS mp)
     | _                             => [thm]
 in map zero_var_indexes (at thm) end;
 
@@ -71,15 +71,15 @@
                                   resolve_tac (maps atomizeD prems) 1);
   -- is nice, but cannot instantiate unknowns in the assumptions *)
 fun hyp_resolve_tac i st = let
-        fun ap (Const("All",_)$Abs(_,_,t))=(case ap t of (k,a,t) => (k+1,a  ,t))
-        |   ap (Const("op -->",_)$_$t)    =(case ap t of (k,_,t) => (k,true ,t))
+        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 t                          =                         (0,false,t);
 (*
         fun rep_goal (Const ("all",_)$Abs (_,_,t)) = rep_goal t
         |   rep_goal (Const ("==>",_)$s$t)         =
                         (case rep_goal t of (l,t) => (s::l,t))
         |   rep_goal t                             = ([]  ,t);
-        val (prems, Const("Trueprop", _)$concl) = rep_goal
+        val (prems, Const(@{const_name Trueprop}, _)$concl) = rep_goal
                                                 (#3(dest_state (st,i)));
 *)
         val subgoal = #3(Thm.dest_state (st,i));
--- a/src/HOL/Recdef.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Recdef.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -5,7 +5,7 @@
 header {* TFL: recursive function definitions *}
 
 theory Recdef
-imports FunDef Plain
+imports Plain Hilbert_Choice
 uses
   ("Tools/TFL/casesplit.ML")
   ("Tools/TFL/utils.ML")
--- a/src/HOL/Record.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Record.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -10,7 +10,7 @@
 
 theory Record
 imports Plain Quickcheck
-uses ("Tools/typecopy.ML") ("Tools/record.ML") ("Tools/quickcheck_record.ML")
+uses ("Tools/record.ML")
 begin
 
 subsection {* Introduction *}
@@ -452,9 +452,7 @@
 
 subsection {* Record package *}
 
-use "Tools/typecopy.ML" setup Typecopy.setup
 use "Tools/record.ML" setup Record.setup
-use "Tools/quickcheck_record.ML" setup Quickcheck_Record.setup
 
 hide_const (open) Tuple_Isomorphism repr abst iso_tuple_fst iso_tuple_snd
   iso_tuple_fst_update iso_tuple_snd_update iso_tuple_cons
--- a/src/HOL/Statespace/distinct_tree_prover.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Statespace/distinct_tree_prover.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -343,7 +343,7 @@
   end handle Option => NONE)
 
 fun distinctTree_tac names ctxt 
-      (Const ("Trueprop",_) $ (Const ("Not", _) $ (Const ("op =", _) $ x $ y)), i) =
+      (Const (@{const_name Trueprop},_) $ (Const (@{const_name Not}, _) $ (Const (@{const_name "op ="}, _) $ 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 @{theory HOL} "DistinctTreeProver.distinct_simproc" ["x = y"]
-    (fn thy => fn ss => fn (Const ("op =",_)$x$y) =>
+    (fn thy => fn ss => fn (Const (@{const_name "op ="},_)$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 19 17:40:44 2010 +0200
+++ b/src/HOL/Statespace/state_fun.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -43,9 +43,9 @@
 val conj_True = thm "conj_True";
 val conj_cong = thm "conj_cong";
 
-fun isFalse (Const ("False",_)) = true
+fun isFalse (Const (@{const_name False},_)) = true
   | isFalse _ = false;
-fun isTrue (Const ("True",_)) = true
+fun isTrue (Const (@{const_name True},_)) = true
   | isTrue _ = false;
 
 in
@@ -53,7 +53,7 @@
 val lazy_conj_simproc =
   Simplifier.simproc @{theory HOL} "lazy_conj_simp" ["P & Q"]
     (fn thy => fn ss => fn t =>
-      (case t of (Const ("op &",_)$P$Q) => 
+      (case t of (Const (@{const_name "op &"},_)$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 ("op =",Teq)$sel'$x',hd (binder_types Teq),nT,swap) end;
+                  in (Const (@{const_name "op ="},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,20 +295,20 @@
            | dest_state s = 
                     raise TERM ("StateFun.ex_lookup_eq_simproc: not a record slector",[s]);
   
-         fun dest_sel_eq (Const ("op =",Teq)$
+         fun dest_sel_eq (Const (@{const_name "op ="},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 ("op =",Teq)$X$
+           | dest_sel_eq (Const (@{const_name "op ="},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 ("",[]);
 
        in
          (case t of
-           (Const ("Ex",Tex)$Abs(s,T,t)) =>
+           (Const (@{const_name Ex},Tex)$Abs(s,T,t)) =>
              (let val (eq,eT,nT,swap) = mkeq (dest_sel_eq t) 0;
                   val prop = list_all ([("n",nT),("x",eT)],
-                              Logic.mk_equals (Const ("Ex",Tex)$Abs(s,T,eq),
+                              Logic.mk_equals (Const (@{const_name Ex},Tex)$Abs(s,T,eq),
                                                HOLogic.true_const));
                   val thm = Drule.export_without_context (prove prop);
                   val thm' = if swap then swap_ex_eq OF [thm] else thm
@@ -367,7 +367,7 @@
      val (lookup_ss,ex_lookup_ss,simprocs_active) = StateFunData.get ctxt;
      val (lookup_ss', ex_lookup_ss') = 
            (case (concl_of thm) of
-            (_$((Const ("Ex",_)$_))) => (lookup_ss, ex_lookup_ss addsimps [thm])
+            (_$((Const (@{const_name Ex},_)$_))) => (lookup_ss, ex_lookup_ss addsimps [thm])
             | _ => (lookup_ss addsimps [thm], ex_lookup_ss))
      fun activate_simprocs ctxt =
           if simprocs_active then ctxt
--- a/src/HOL/Statespace/state_space.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Statespace/state_space.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -222,8 +222,8 @@
   end handle Option => NONE)
 
 fun distinctTree_tac ctxt
-      (Const ("Trueprop",_) $
-        (Const ("Not", _) $ (Const ("op =", _) $ (x as Free _)$ (y as Free _))), i) =
+      (Const (@{const_name Trueprop},_) $
+        (Const (@{const_name Not}, _) $ (Const (@{const_name "op ="}, _) $ (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 @{theory HOL} "StateSpace.distinct_simproc" ["x = y"]
-    (fn thy => fn ss => (fn (Const ("op =",_)$(x as Free _)$(y as Free _)) =>
+    (fn thy => fn ss => (fn (Const (@{const_name "op ="},_)$(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)
--- a/src/HOL/TLA/Intensional.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/TLA/Intensional.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -279,7 +279,7 @@
 
     fun hflatten t =
         case (concl_of t) of
-          Const _ $ (Const ("op -->", _) $ _ $ _) => hflatten (t RS mp)
+          Const _ $ (Const (@{const_name "op -->"}, _) $ _ $ _) => hflatten (t RS mp)
         | _ => (hflatten (matchsome conjI t)) handle THM _ => zero_var_indexes t
   in
     hflatten t
--- a/src/HOL/Tools/ATP/atp_problem.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/ATP/atp_problem.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -63,12 +63,12 @@
   | string_for_connective AIff = "<=>"
   | string_for_connective ANotIff = "<~>"
 fun string_for_formula (AQuant (q, xs, phi)) =
-    string_for_quantifier q ^ "[" ^ commas xs ^ "] : " ^
-    string_for_formula phi
+    "(" ^ string_for_quantifier q ^ "[" ^ commas xs ^ "] : " ^
+    string_for_formula phi ^ ")"
   | string_for_formula (AConn (ANot, [AAtom (ATerm ("equal", ts))])) =
     space_implode " != " (map string_for_term ts)
   | string_for_formula (AConn (c, [phi])) =
-    string_for_connective c ^ " " ^ string_for_formula phi
+    "(" ^ string_for_connective c ^ " " ^ string_for_formula phi ^ ")"
   | string_for_formula (AConn (c, phis)) =
     "(" ^ space_implode (" " ^ string_for_connective c ^ " ")
                         (map string_for_formula phis) ^ ")"
--- a/src/HOL/Tools/ATP/atp_systems.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/ATP/atp_systems.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -8,9 +8,9 @@
 signature ATP_SYSTEMS =
 sig
   datatype failure =
-    Unprovable | IncompleteUnprovable | CantConnect | TimedOut | OutOfResources |
-    OldSpass | OldVampire | NoPerl | NoLibwwwPerl | MalformedInput |
-    MalformedOutput | UnknownError
+    Unprovable | IncompleteUnprovable | CantConnect | TimedOut |
+    OutOfResources | SpassTooOld | VampireTooOld | NoPerl | NoLibwwwPerl |
+    MalformedInput | MalformedOutput | UnknownError
 
   type prover_config =
     {exec: string * string,
@@ -40,7 +40,7 @@
 
 datatype failure =
   Unprovable | IncompleteUnprovable | CantConnect | TimedOut | OutOfResources |
-  OldSpass | OldVampire | NoPerl | NoLibwwwPerl | MalformedInput |
+  SpassTooOld | VampireTooOld | NoPerl | NoLibwwwPerl | MalformedInput |
   MalformedOutput | UnknownError
 
 type prover_config =
@@ -63,7 +63,7 @@
   | string_for_failure CantConnect = "Can't connect to remote server."
   | string_for_failure TimedOut = "Timed out."
   | string_for_failure OutOfResources = "The ATP ran out of resources."
-  | string_for_failure OldSpass =
+  | string_for_failure SpassTooOld =
     "Isabelle requires a more recent version of SPASS with support for the \
     \TPTP syntax. To install it, download and extract the package \
     \\"http://isabelle.in.tum.de/dist/contrib/spass-3.7.tar.gz\" and add the \
@@ -72,7 +72,7 @@
                (Path.variable "ISABELLE_HOME_USER" ::
                 map Path.basic ["etc", "components"])))) ^
     " on a line of its own."
-  | string_for_failure OldVampire =
+  | string_for_failure VampireTooOld =
     "Isabelle requires a more recent version of Vampire. To install it, follow \
     \the instructions from the Sledgehammer manual (\"isabelle doc\
     \ sledgehammer\")."
@@ -144,7 +144,7 @@
        "# Cannot determine problem status within resource limit"),
       (OutOfResources, "SZS status: ResourceOut"),
       (OutOfResources, "SZS status ResourceOut")],
-   max_new_relevant_facts_per_iter = 60 (* FIXME *),
+   max_new_relevant_facts_per_iter = 50 (* FIXME *),
    prefers_theory_relevant = false,
    explicit_forall = false}
 
@@ -170,7 +170,7 @@
       (OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"),
       (MalformedInput, "Undefined symbol"),
       (MalformedInput, "Free Variable"),
-      (OldSpass, "tptp2dfg")],
+      (SpassTooOld, "tptp2dfg")],
    max_new_relevant_facts_per_iter = 35 (* FIXME *),
    prefers_theory_relevant = true,
    explicit_forall = true}
@@ -185,7 +185,7 @@
    required_execs = [],
    arguments = fn _ => fn timeout =>
      "--mode casc -t " ^ string_of_int (to_generous_secs timeout) ^
-     " --input_file ",
+     " --input_file",
    proof_delims =
      [("=========== Refutation ==========",
        "======= End of refutation ======="),
@@ -197,8 +197,8 @@
       (TimedOut, "SZS status Timeout"),
       (Unprovable, "Satisfiability detected"),
       (OutOfResources, "Refutation not found"),
-      (OldVampire, "input_file")],
-   max_new_relevant_facts_per_iter = 50 (* FIXME *),
+      (VampireTooOld, "not a valid option")],
+   max_new_relevant_facts_per_iter = 45 (* FIXME *),
    prefers_theory_relevant = false,
    explicit_forall = false}
 
--- a/src/HOL/Tools/Datatype/datatype_codegen.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_codegen.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -16,9 +16,9 @@
 
 (* liberal addition of code data for datatypes *)
 
-fun mk_constr_consts thy vs dtco cos =
+fun mk_constr_consts thy vs tyco cos =
   let
-    val cs = map (fn (c, tys) => (c, tys ---> Type (dtco, map TFree vs))) cos;
+    val cs = map (fn (c, tys) => (c, tys ---> Type (tyco, map TFree vs))) cos;
     val cs' = map (fn c_ty as (_, ty) => (AxClass.unoverload_const thy c_ty, ty)) cs;
   in if is_some (try (Code.constrset_of_consts thy) cs')
     then SOME cs
@@ -62,12 +62,12 @@
 
 (* equality *)
 
-fun mk_eq_eqns thy dtco =
+fun mk_eq_eqns thy tyco =
   let
-    val (vs, cos) = Datatype_Data.the_spec thy dtco;
+    val (vs, cos) = Datatype_Data.the_spec thy tyco;
     val { descr, index, inject = inject_thms, distinct = distinct_thms, ... } =
-      Datatype_Data.the_info thy dtco;
-    val ty = Type (dtco, map TFree vs);
+      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)
       $ t1 $ t2;
     fun true_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.true_const);
@@ -88,11 +88,11 @@
       |> Simpdata.mk_eq;
   in (map prove (triv_injects @ injects @ distincts), prove refl) end;
 
-fun add_equality vs dtcos thy =
+fun add_equality vs tycos thy =
   let
-    fun add_def dtco lthy =
+    fun add_def tyco lthy =
       let
-        val ty = Type (dtco, map TFree vs);
+        val ty = Type (tyco, map TFree vs);
         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
@@ -105,35 +105,37 @@
       in (thm', lthy') end;
     fun tac thms = Class.intro_classes_tac []
       THEN ALLGOALS (ProofContext.fact_tac thms);
-    fun prefix dtco = Binding.qualify true (Long_Name.base_name dtco) o Binding.qualify true "eq" o Binding.name;
-    fun add_eq_thms dtco =
+    fun prefix tyco = Binding.qualify true (Long_Name.base_name tyco) o Binding.qualify true "eq" o Binding.name;
+    fun add_eq_thms tyco =
       Theory.checkpoint
-      #> `(fn thy => mk_eq_eqns thy dtco)
+      #> `(fn thy => mk_eq_eqns thy tyco)
       #-> (fn (thms, thm) => PureThy.note_thmss Thm.lemmaK
-          [((prefix dtco "refl", [Code.add_nbe_default_eqn_attribute]), [([thm], [])]),
-            ((prefix dtco "simps", [Code.add_default_eqn_attribute]), [(rev thms, [])])])
+          [((prefix tyco "refl", [Code.add_nbe_default_eqn_attribute]), [([thm], [])]),
+            ((prefix tyco "simps", [Code.add_default_eqn_attribute]), [(rev thms, [])])])
       #> snd
   in
     thy
-    |> Class.instantiation (dtcos, vs, [HOLogic.class_eq])
-    |> fold_map add_def dtcos
+    |> Class.instantiation (tycos, vs, [HOLogic.class_eq])
+    |> 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)
     |-> (fn def_thms => fold Code.del_eqn def_thms)
-    |> fold add_eq_thms dtcos
+    |> fold add_eq_thms tycos
   end;
 
 
 (* register a datatype etc. *)
 
-fun add_all_code config dtcos thy =
+fun add_all_code config tycos thy =
   let
-    val (vs :: _, coss) = (split_list o map (Datatype_Data.the_spec thy)) dtcos;
-    val any_css = map2 (mk_constr_consts thy vs) dtcos coss;
+    val (vs :: _, coss) = (split_list o map (Datatype_Data.the_spec thy)) tycos;
+    val any_css = map2 (mk_constr_consts thy vs) tycos coss;
     val css = if exists is_none any_css then []
       else map_filter I any_css;
-    val case_rewrites = maps (#case_rewrites o Datatype_Data.the_info thy) dtcos;
-    val certs = map (mk_case_cert thy) dtcos;
+    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;
   in
     if null css then thy
     else thy
@@ -141,7 +143,7 @@
       |> fold Code.add_datatype css
       |> fold_rev Code.add_default_eqn case_rewrites
       |> fold Code.add_case certs
-      |> add_equality vs dtcos
+      |> not (null tycos_eq) ? add_equality vs tycos_eq
    end;
 
 
--- a/src/HOL/Tools/Function/function_core.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Function/function_core.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -392,7 +392,7 @@
     (* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *)
     val ihyp = Term.all domT $ Abs ("z", domT,
       Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
-        HOLogic.mk_Trueprop (Const ("Ex1", (ranT --> boolT) --> boolT) $
+        HOLogic.mk_Trueprop (Const (@{const_name Ex1}, (ranT --> boolT) --> boolT) $
           Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
       |> cterm_of thy
 
--- a/src/HOL/Tools/Function/termination.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Function/termination.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -148,8 +148,8 @@
     val cs = Function_Lib.dest_binop_list @{const_name Lattices.sup} rel
     fun collect_pats (Const (@{const_name Collect}, _) $ Abs (_, _, c)) =
       let
-        val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ _)
-          = Term.strip_qnt_body "Ex" c
+        val (Const (@{const_name "op &"}, _) $ (Const (@{const_name "op ="}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ _)
+          = Term.strip_qnt_body @{const_name Ex} c
       in cons r o cons l end
   in
     mk_skel (fold collect_pats cs [])
@@ -182,11 +182,11 @@
 fun dest_call D (Const (@{const_name Collect}, _) $ Abs (_, _, c)) =
   let
     val (sk, _, _, _, _) = D
-    val vs = Term.strip_qnt_vars "Ex" c
+    val vs = Term.strip_qnt_vars @{const_name Ex} c
 
     (* FIXME: throw error "dest_call" for malformed terms *)
-    val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ Gam)
-      = Term.strip_qnt_body "Ex" c
+    val (Const (@{const_name "op &"}, _) $ (Const (@{const_name "op ="}, _) $ _ $ (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
   in
--- a/src/HOL/Tools/Nitpick/kodkod.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Nitpick/kodkod.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -156,6 +156,7 @@
 
   datatype outcome =
     JavaNotInstalled |
+    JavaTooOld |
     KodkodiNotInstalled |
     Normal of (int * raw_bound list) list * int list * string |
     TimedOut of int list |
@@ -303,6 +304,7 @@
 
 datatype outcome =
   JavaNotInstalled |
+  JavaTooOld |
   KodkodiNotInstalled |
   Normal of (int * raw_bound list) list * int list * string |
   TimedOut of int list |
@@ -1063,19 +1065,20 @@
                 |> perhaps (try (unsuffix "."))
                 |> perhaps (try (unprefix "Solve error: "))
                 |> perhaps (try (unprefix "Error: "))
+              fun has_error s =
+                first_error |> Substring.full |> Substring.position s |> snd
+                            |> Substring.isEmpty |> not
             in
               if null ps then
                 if code = 2 then
                   TimedOut js
                 else if code = 0 then
                   Normal ([], js, first_error)
-                else if first_error |> Substring.full
-                        |> Substring.position "exec: java" |> snd
-                        |> Substring.isEmpty |> not then
+                else if has_error "exec: java" then
                   JavaNotInstalled
-                else if first_error |> Substring.full
-                        |> Substring.position "NoClassDefFoundError" |> snd
-                        |> Substring.isEmpty |> not then
+                else if has_error "UnsupportedClassVersionError" then
+                  JavaTooOld
+                else if has_error "NoClassDefFoundError" then
                   KodkodiNotInstalled
                 else if first_error <> "" then
                   Error (first_error, js)
--- a/src/HOL/Tools/Nitpick/minipick.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Nitpick/minipick.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -315,6 +315,7 @@
   in
     case solve_any_problem overlord NONE max_threads max_solutions problems of
       JavaNotInstalled => "unknown"
+    | JavaTooOld => "unknown"
     | KodkodiNotInstalled => "unknown"
     | Normal ([], _, _) => "none"
     | Normal _ => "genuine"
--- a/src/HOL/Tools/Nitpick/nitpick.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -740,6 +740,9 @@
             KK.JavaNotInstalled =>
             (print_m install_java_message;
              (found_really_genuine, max_potential, max_genuine, donno + 1))
+          | KK.JavaTooOld =>
+            (print_m install_java_message;
+             (found_really_genuine, max_potential, max_genuine, donno + 1))
           | KK.KodkodiNotInstalled =>
             (print_m install_kodkodi_message;
              (found_really_genuine, max_potential, max_genuine, donno + 1))
--- a/src/HOL/Tools/Predicate_Compile/code_prolog.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/code_prolog.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -147,7 +147,7 @@
 
 fun Trueprop_conv cv ct =
   case Thm.term_of ct of
-    Const ("Trueprop", _) $ _ => Conv.arg_conv cv ct  
+    Const (@{const_name Trueprop}, _) $ _ => Conv.arg_conv cv ct  
   | _ => raise Fail "Trueprop_conv"
 
 fun preprocess_intro thy rule =
@@ -414,6 +414,6 @@
 val _ = Outer_Syntax.improper_command "values" "enumerate and print comprehensions" Keyword.diag
   (opt_print_modes -- Scan.optional (Parse.nat >> SOME) NONE -- Parse.term
    >> (fn ((print_modes, soln), t) => Toplevel.keep
-        (values_cmd print_modes soln t)));
+        (values_cmd print_modes soln t))); (*FIXME does not preserve the previous functionality*)
 
 end;
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -405,13 +405,13 @@
 (* general syntactic functions *)
 
 (*Like dest_conj, but flattens conjunctions however nested*)
-fun conjuncts_aux (Const ("op &", _) $ t $ t') conjs = conjuncts_aux t (conjuncts_aux t' conjs)
+fun conjuncts_aux (Const (@{const_name "op &"}, _) $ 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 ("Trueprop", _) $ (Const ("op =", _) $ _ $ _)) = true
+  | is_equationlike_term (Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, _) $ _ $ _)) = true
   | is_equationlike_term _ = false
   
 val is_equationlike = is_equationlike_term o prop_of 
@@ -429,10 +429,7 @@
   
 fun is_intro constname t = is_intro_term constname (prop_of t)
 
-fun is_pred thy constname =
-  let
-    val T = (Sign.the_const_type thy constname)
-  in body_type T = @{typ "bool"} end;
+fun is_pred thy constname = (body_type (Sign.the_const_type thy constname) = HOLogic.boolT);
 
 fun is_predT (T as Type("fun", [_, _])) = (snd (strip_type T) = @{typ bool})
   | is_predT _ = false
@@ -482,7 +479,7 @@
 
 fun strip_all t = (Term.strip_all_vars t, Term.strip_all_body t)
 
-fun strip_ex (Const ("Ex", _) $ Abs (x, T, t)) =
+fun strip_ex (Const (@{const_name Ex}, _) $ Abs (x, T, t)) =
   let
     val (xTs, t') = strip_ex t
   in
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -524,7 +524,7 @@
 
 fun dest_conjunct_prem th =
   case HOLogic.dest_Trueprop (prop_of th) of
-    (Const ("op &", _) $ t $ t') =>
+    (Const (@{const_name "op &"}, _) $ t $ t') =>
       dest_conjunct_prem (th RS @{thm conjunct1})
         @ dest_conjunct_prem (th RS @{thm conjunct2})
     | _ => [th]
@@ -576,7 +576,7 @@
 
 fun Trueprop_conv cv ct =
   case Thm.term_of ct of
-    Const ("Trueprop", _) $ _ => Conv.arg_conv cv ct  
+    Const (@{const_name Trueprop}, _) $ _ => Conv.arg_conv cv ct  
   | _ => raise Fail "Trueprop_conv"
 
 fun preprocess_intro thy rule =
@@ -587,7 +587,7 @@
 
 fun preprocess_elim ctxt elimrule =
   let
-    fun replace_eqs (Const ("Trueprop", _) $ (Const ("op =", T) $ lhs $ rhs)) =
+    fun replace_eqs (Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, T) $ lhs $ rhs)) =
        HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs)
      | replace_eqs t = t
     val thy = ProofContext.theory_of ctxt
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -111,7 +111,7 @@
 
 fun mk_meta_equation th =
   case prop_of th of
-    Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ _) => th RS @{thm eq_reflection}
+    Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, _) $ _ $ _) => th RS @{thm eq_reflection}
   | _ => th
 
 val meta_fun_cong = @{lemma "f == g ==> f x == g x" by simp}
@@ -215,12 +215,12 @@
 val logic_operator_names =
   [@{const_name "=="}, 
    @{const_name "==>"},
-   @{const_name "Trueprop"},
-   @{const_name "Not"},
+   @{const_name Trueprop},
+   @{const_name Not},
    @{const_name "op ="},
    @{const_name "op -->"},
-   @{const_name "All"},
-   @{const_name "Ex"}, 
+   @{const_name All},
+   @{const_name Ex}, 
    @{const_name "op &"},
    @{const_name "op |"}]
 
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -52,7 +52,7 @@
 fun transform_ho_typ (T as Type ("fun", _)) =
   let
     val (Ts, T') = strip_type T
-  in if T' = @{typ "bool"} then T else (Ts @ [T']) ---> HOLogic.boolT end
+  in if T' = HOLogic.boolT then T else (Ts @ [T']) ---> HOLogic.boolT end
 | transform_ho_typ t = t
 
 fun transform_ho_arg arg = 
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -86,11 +86,11 @@
    map instantiate rew_ths
  end
 
-fun is_compound ((Const ("Not", _)) $ t) =
+fun is_compound ((Const (@{const_name Not}, _)) $ t) =
     error "is_compound: Negation should not occur; preprocessing is defect"
-  | is_compound ((Const ("Ex", _)) $ _) = true
+  | is_compound ((Const (@{const_name Ex}, _)) $ _) = true
   | is_compound ((Const (@{const_name "op |"}, _)) $ _ $ _) = true
-  | is_compound ((Const ("op &", _)) $ _ $ _) =
+  | is_compound ((Const (@{const_name "op &"}, _)) $ _ $ _) =
     error "is_compound: Conjunction should not occur; preprocessing is defect"
   | is_compound _ = false
 
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -168,10 +168,10 @@
     mk_split_lambda' xs t
   end;
 
-fun strip_imp_prems (Const("op -->", _) $ A $ B) = A :: strip_imp_prems B
+fun strip_imp_prems (Const(@{const_name "op -->"}, _) $ A $ B) = A :: strip_imp_prems B
   | strip_imp_prems _ = [];
 
-fun strip_imp_concl (Const("op -->", _) $ A $ B) = strip_imp_concl B
+fun strip_imp_concl (Const(@{const_name "op -->"}, _) $ A $ B) = strip_imp_concl B
   | strip_imp_concl A = A : term;
 
 fun strip_horn A = (strip_imp_prems A, strip_imp_concl A);
--- a/src/HOL/Tools/Qelim/cooper.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Qelim/cooper.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -120,10 +120,10 @@
 
 fun whatis x ct =
 ( case (term_of ct) of
-  Const("op &",_)$_$_ => And (Thm.dest_binop ct)
-| Const ("op |",_)$_$_ => Or (Thm.dest_binop ct)
-| Const ("op =",_)$y$_ => if term_of x aconv y then Eq (Thm.dest_arg ct) else Nox
-| Const (@{const_name Not},_) $ (Const ("op =",_)$y$_) =>
+  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$_) =>
   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("op =",_))$s$t) =
+  | lin (vs as x::_) ((b as Const(@{const_name "op ="},_))$s$t) =
      (case lint vs (subC$t$s) of
       (t as a$(m$c$y)$r) =>
         if x <> y then b$zero$t
@@ -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("op &",_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
-  | Const("op |",_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
+  | 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 Not},_)$_ => h (acc,dacc) (Thm.dest_arg t)
   | _ => (acc, dacc)
   val (cs,ds) = h ([],[]) p
@@ -382,8 +382,8 @@
     end
   fun unit_conv t =
    case (term_of t) of
-   Const("op &",_)$_$_ => Conv.binop_conv unit_conv t
-  | Const("op |",_)$_$_ => Conv.binop_conv unit_conv t
+   Const(@{const_name "op &"},_)$_$_ => Conv.binop_conv unit_conv t
+  | Const(@{const_name "op |"},_)$_$_ => 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 =)
@@ -612,19 +612,19 @@
 
 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 ("op &", _) $ t1 $ t2) =
+  | fm_of_term ps vs (Const (@{const_name "op &"}, _) $ t1 $ t2) =
       Proc.And (fm_of_term ps vs t1, fm_of_term ps vs t2)
-  | fm_of_term ps vs (Const ("op |", _) $ t1 $ t2) =
+  | fm_of_term ps vs (Const (@{const_name "op |"}, _) $ t1 $ t2) =
       Proc.Or (fm_of_term ps vs t1, fm_of_term ps vs t2)
-  | fm_of_term ps vs (Const ("op -->", _) $ t1 $ t2) =
+  | fm_of_term ps vs (Const (@{const_name "op -->"}, _) $ 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)
   | fm_of_term ps vs (Const (@{const_name Not}, _) $ t') =
       Proc.Not (fm_of_term ps vs t')
-  | fm_of_term ps vs (Const ("Ex", _) $ Abs abs) =
+  | fm_of_term ps vs (Const (@{const_name Ex}, _) $ Abs abs) =
       Proc.E (uncurry (fm_of_term ps) (descend vs abs))
-  | fm_of_term ps vs (Const ("All", _) $ Abs abs) =
+  | fm_of_term ps vs (Const (@{const_name All}, _) $ Abs abs) =
       Proc.A (uncurry (fm_of_term ps) (descend vs abs))
   | fm_of_term ps vs (@{term "op = :: int => _"} $ t1 $ t2) =
       Proc.Eq (Proc.Sub (num_of_term vs t1, num_of_term vs t2))
@@ -687,14 +687,14 @@
 
 fun strip_objimp ct =
   (case Thm.term_of ct of
-    Const ("op -->", _) $ _ $ _ =>
+    Const (@{const_name "op -->"}, _) $ _ $ _ =>
       let val (A, B) = Thm.dest_binop ct
       in A :: strip_objimp B end
   | _ => [ct]);
 
 fun strip_objall ct = 
  case term_of ct of 
-  Const ("All", _) $ Abs (xn,xT,p) => 
+  Const (@{const_name All}, _) $ Abs (xn,xT,p) => 
    let val (a,(v,t')) = (apsnd (Thm.dest_abs (SOME xn)) o Thm.dest_comb) ct
    in apfst (cons (a,v)) (strip_objall t')
    end
--- a/src/HOL/Tools/Quotient/quotient_term.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Quotient/quotient_term.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -490,16 +490,16 @@
          else mk_bex1_rel $ resrel $ subtrm
        end
 
-  | (Const (@{const_name "Ex1"}, ty) $ t, Const (@{const_name "Ex1"}, ty') $ t') =>
+  | (Const (@{const_name Ex1}, ty) $ t, Const (@{const_name Ex1}, ty') $ t') =>
        let
          val subtrm = apply_subt (regularize_trm ctxt) (t, t')
        in
-         if ty = ty' then Const (@{const_name "Ex1"}, ty) $ subtrm
+         if ty = ty' then Const (@{const_name Ex1}, ty) $ subtrm
          else mk_bex1_rel $ (equiv_relation ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm
        end
 
   | (Const (@{const_name "Ball"}, ty) $ (Const (@{const_name "Respects"}, _) $ resrel) $ t,
-     Const (@{const_name "All"}, ty') $ t') =>
+     Const (@{const_name All}, ty') $ t') =>
        let
          val subtrm = apply_subt (regularize_trm ctxt) (t, t')
          val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty')
@@ -510,7 +510,7 @@
        end
 
   | (Const (@{const_name "Bex"}, ty) $ (Const (@{const_name "Respects"}, _) $ resrel) $ t,
-     Const (@{const_name "Ex"}, ty') $ t') =>
+     Const (@{const_name Ex}, ty') $ t') =>
        let
          val subtrm = apply_subt (regularize_trm ctxt) (t, t')
          val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty')
@@ -520,7 +520,7 @@
          else mk_bex $ (mk_resp $ resrel) $ subtrm
        end
 
-  | (Const (@{const_name "Bex1_rel"}, ty) $ resrel $ t, Const (@{const_name "Ex1"}, ty') $ t') =>
+  | (Const (@{const_name "Bex1_rel"}, ty) $ resrel $ t, Const (@{const_name Ex1}, ty') $ t') =>
        let
          val subtrm = apply_subt (regularize_trm ctxt) (t, t')
          val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty')
@@ -638,10 +638,10 @@
    as the type of subterms needs to be calculated   *)
 fun inj_repabs_trm ctxt (rtrm, qtrm) =
  case (rtrm, qtrm) of
-    (Const (@{const_name "Ball"}, T) $ r $ t, Const (@{const_name "All"}, _) $ t') =>
+    (Const (@{const_name "Ball"}, T) $ r $ t, Const (@{const_name All}, _) $ t') =>
        Const (@{const_name "Ball"}, T) $ r $ (inj_repabs_trm ctxt (t, t'))
 
-  | (Const (@{const_name "Bex"}, T) $ r $ t, Const (@{const_name "Ex"}, _) $ t') =>
+  | (Const (@{const_name "Bex"}, T) $ r $ t, Const (@{const_name Ex}, _) $ t') =>
        Const (@{const_name "Bex"}, T) $ r $ (inj_repabs_trm ctxt (t, t'))
 
   | (Const (@{const_name "Babs"}, T) $ r $ t, t' as (Abs _)) =>
--- a/src/HOL/Tools/Sledgehammer/metis_tactics.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/metis_tactics.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -308,7 +308,7 @@
                   SOME w =>  mk_var(w, dummyT)
                 | NONE   => mk_var(v, dummyT))
         | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn ("=",[]), _])) =
-            Const ("op =", HOLogic.typeT)
+            Const (@{const_name "op ="}, HOLogic.typeT)
         | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (x,[]), ty])) =
            (case strip_prefix_and_undo_ascii const_prefix x of
                 SOME c => Const (invert_const c, dummyT)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -167,8 +167,8 @@
 val set_ClauseFormulaRelationN = "set_ClauseFormulaRelation"
 
 val parse_clause_formula_pair =
-  $$ "(" |-- scan_integer --| $$ "," -- Symbol.scan_id
-  --| Scan.repeat ($$ "," |-- Symbol.scan_id) --| $$ ")"
+  $$ "(" |-- scan_integer --| $$ ","
+  -- (Symbol.scan_id ::: Scan.repeat ($$ "," |-- Symbol.scan_id)) --| $$ ")"
   --| Scan.option ($$ ",")
 val parse_clause_formula_relation =
   Scan.this_string set_ClauseFormulaRelationN |-- $$ "("
@@ -183,21 +183,21 @@
                                               thm_names =
   if String.isSubstring set_ClauseFormulaRelationN output then
     (* This is a hack required for keeping track of axioms after they have been
-       clausified by SPASS's Flotter tool. The "SPASS_TPTP" script is also part
-       of this hack. *)
+       clausified by SPASS's Flotter tool. The "ATP/scripts/spass" script is also
+       part of this hack. *)
     let
       val j0 = hd (hd conjecture_shape)
       val seq = extract_clause_sequence output
       val name_map = extract_clause_formula_relation output
       fun renumber_conjecture j =
-        AList.find (op =) name_map (conjecture_prefix ^ Int.toString (j - j0))
+        conjecture_prefix ^ Int.toString (j - j0)
+        |> AList.find (fn (s, ss) => member (op =) ss s) name_map
         |> map (fn s => find_index (curry (op =) s) seq + 1)
     in
       (conjecture_shape |> map (maps renumber_conjecture),
-       seq |> map (fn j => case j |> AList.lookup (op =) name_map |> the
-                                  |> try (unprefix axiom_prefix) of
-                             SOME s' => undo_ascii_of s'
-                           | NONE => "")
+       seq |> map (AList.lookup (op =) name_map #> these
+                   #> map_filter (try (unprefix axiom_prefix))
+                   #> map undo_ascii_of #> space_implode " ")
            |> Vector.fromList)
     end
   else
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimize.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimize.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -85,8 +85,7 @@
 fun sublinear_minimize _ [] p = p
   | sublinear_minimize test (x :: xs) (seen, result) =
     case test (xs @ seen) of
-      result as {outcome = NONE, proof, used_thm_names, ...}
-      : prover_result =>
+      result as {outcome = NONE, proof, used_thm_names, ...} : prover_result =>
       sublinear_minimize test (filter_used_facts used_thm_names xs)
                          (filter_used_facts used_thm_names seen, result)
     | _ => sublinear_minimize test xs (x :: seen, result)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -279,10 +279,10 @@
   | add_type_constraint (false, cl, TVar (ix, _)) = add_var (ix, cl)
   | add_type_constraint _ = I
 
-fun fix_atp_variable_name s =
+fun repair_atp_variable_name f s =
   let
     fun subscript_name s n = s ^ nat_subscript n
-    val s = String.map Char.toLower s
+    val s = String.map f s
   in
     case space_explode "_" s of
       [_] => (case take_suffix Char.isDigit (String.explode s) of
@@ -349,9 +349,10 @@
                   SOME b => Var ((b, 0), T)
                 | NONE =>
                   if is_tptp_variable a then
-                    Var ((fix_atp_variable_name a, 0), T)
+                    Var ((repair_atp_variable_name Char.toLower a, 0), T)
                   else
-                    raise Fail ("Unexpected constant: " ^ quote a)
+                    (* Skolem constants? *)
+                    Var ((repair_atp_variable_name Char.toUpper a, 0), T)
           in list_comb (t, ts) end
   in aux (SOME HOLogic.boolT) [] end
 
@@ -410,7 +411,8 @@
         do_formula pos (AQuant (q, xs, phi'))
         #>> quantify_over_free (case q of
                                   AForall => @{const_name All}
-                                | AExists => @{const_name Ex}) x
+                                | AExists => @{const_name Ex})
+                               (repair_atp_variable_name Char.toLower x)
       | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not
       | AConn (c, [phi1, phi2]) =>
         do_formula (pos |> c = AImplies ? not) phi1
@@ -834,6 +836,7 @@
       second_pass [Show] (contra_proof, [], ([], ([], []))) |> fst
   in proof_top @ proof_bottom end
 
+(* FIXME: Still needed? Probably not. *)
 val kill_duplicate_assumptions_in_proof =
   let
     fun relabel_facts subst =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -55,8 +55,10 @@
   let
     val do_term = combterm_from_term thy
     fun do_quant bs q s T t' =
-      do_formula ((s, T) :: bs) t'
-      #>> (fn phi => AQuant (q, [`make_bound_var s], phi))
+      let val s = Name.variant (map fst bs) s in
+        do_formula ((s, T) :: bs) t'
+        #>> (fn phi => AQuant (q, [`make_bound_var s], phi))
+      end
     and do_conn bs c t1 t2 =
       do_formula bs t1 ##>> do_formula bs t2
       #>> (fn (phi1, phi2) => AConn (c, [phi1, phi2]))
@@ -102,41 +104,41 @@
                     (0 upto length Ts - 1 ~~ Ts))
 
 fun introduce_combinators_in_term ctxt kind t =
-  let
-    val thy = ProofContext.theory_of ctxt
-    fun aux Ts t =
-      case t of
-        @{const Not} $ t1 => @{const Not} $ aux Ts t1
-      | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
-        t0 $ Abs (s, T, aux (T :: Ts) t')
-      | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
-        t0 $ Abs (s, T, aux (T :: Ts) t')
-      | (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}, _])))
-          $ t1 $ t2 =>
-        t0 $ aux Ts t1 $ aux Ts t2
-      | _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then
-               t
-             else
-               let
-                 val t = t |> conceal_bounds Ts
-                           |> Envir.eta_contract
-                 val ([t], ctxt') = Variable.import_terms true [t] ctxt
-               in
-                 t |> cterm_of thy
-                   |> Clausifier.introduce_combinators_in_cterm
-                   |> singleton (Variable.export ctxt' ctxt)
-                   |> prop_of |> Logic.dest_equals |> snd
-                   |> reveal_bounds Ts
-               end
-  in t |> not (Meson.is_fol_term thy t) ? aux [] end
-  handle THM _ =>
-         (* A type variable of sort "{}" will make abstraction fail. *)
-         case kind of
-           Axiom => HOLogic.true_const
-         | Conjecture => HOLogic.false_const
+  let val thy = ProofContext.theory_of ctxt in
+    if Meson.is_fol_term thy t then
+      t
+    else
+      let
+        fun aux Ts t =
+          case t of
+            @{const Not} $ t1 => @{const Not} $ aux Ts t1
+          | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
+            t0 $ Abs (s, T, aux (T :: Ts) t')
+          | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
+            t0 $ Abs (s, T, aux (T :: Ts) t')
+          | (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}, _])))
+              $ t1 $ t2 =>
+            t0 $ aux Ts t1 $ aux Ts t2
+          | _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then
+                   t
+                 else
+                   t |> conceal_bounds Ts
+                     |> Envir.eta_contract
+                     |> cterm_of thy
+                     |> Clausifier.introduce_combinators_in_cterm
+                     |> prop_of |> Logic.dest_equals |> snd
+                     |> reveal_bounds Ts
+        val ([t], ctxt') = Variable.import_terms true [t] ctxt
+      in t |> aux [] |> singleton (Variable.export_terms ctxt' ctxt) end
+      handle THM _ =>
+             (* A type variable of sort "{}" will make abstraction fail. *)
+             case kind of
+               Axiom => HOLogic.true_const
+             | Conjecture => HOLogic.false_const
+  end
 
 (* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the
    same in Sledgehammer to prevent the discovery of unreplable proofs. *)
@@ -267,16 +269,18 @@
         val (x, ty_args) =
           case head of
             CombConst (name as (s, s'), _, ty_args) =>
-            if s = "equal" then
-              (if top_level andalso length args = 2 then name
-               else ("c_fequal", @{const_name fequal}), [])
-            else if top_level then
-              case s of
-                "c_False" => (("$false", s'), [])
-              | "c_True" => (("$true", s'), [])
-              | _ => (name, if full_types then [] else ty_args)
-            else
-              (name, if full_types then [] else ty_args)
+            let val ty_args = if full_types then [] else ty_args in
+              if s = "equal" then
+                if top_level andalso length args = 2 then (name, [])
+                else (("c_fequal", @{const_name fequal}), ty_args)
+              else if top_level then
+                case s of
+                  "c_False" => (("$false", s'), [])
+                | "c_True" => (("$true", s'), [])
+                | _ => (name, ty_args)
+              else
+                (name, ty_args)
+            end
           | CombVar (name, _) => (name, [])
           | CombApp _ => raise Fail "impossible \"CombApp\""
         val t = ATerm (x, map fo_term_for_combtyp ty_args @
--- a/src/HOL/Tools/TFL/dcterm.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/TFL/dcterm.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -24,18 +24,15 @@
   val dest_exists: cterm -> cterm * cterm
   val dest_forall: cterm -> cterm * cterm
   val dest_imp: cterm -> cterm * cterm
-  val dest_let: cterm -> cterm * cterm
   val dest_neg: cterm -> cterm
   val dest_pair: cterm -> cterm * cterm
   val dest_var: cterm -> {Name:string, Ty:typ}
   val is_conj: cterm -> bool
-  val is_cons: cterm -> bool
   val is_disj: cterm -> bool
   val is_eq: cterm -> bool
   val is_exists: cterm -> bool
   val is_forall: cterm -> bool
   val is_imp: cterm -> bool
-  val is_let: cterm -> bool
   val is_neg: cterm -> bool
   val is_pair: cterm -> bool
   val list_mk_disj: cterm list -> cterm
@@ -78,15 +75,15 @@
 
 fun mk_exists (r as (Bvar, Body)) =
   let val ty = #T(rep_cterm Bvar)
-      val c = mk_hol_const("Ex", (ty --> HOLogic.boolT) --> HOLogic.boolT)
+      val c = mk_hol_const(@{const_name Ex}, (ty --> HOLogic.boolT) --> HOLogic.boolT)
   in capply c (uncurry cabs r) end;
 
 
-local val c = mk_hol_const("op &", HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
+local val c = mk_hol_const(@{const_name "op &"}, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
 in fun mk_conj(conj1,conj2) = capply (capply c conj1) conj2
 end;
 
-local val c = mk_hol_const("op |", HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
+local val c = mk_hol_const(@{const_name "op |"}, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
 in fun mk_disj(disj1,disj2) = capply (capply c disj1) disj2
 end;
 
@@ -128,17 +125,15 @@
   handle U.ERR _ => raise ERR "dest_binder" ("Not a(n) " ^ quote expected);
 
 
-val dest_neg    = dest_monop "not"
-val dest_pair   = dest_binop @{const_name Pair};
-val dest_eq     = dest_binop "op ="
-val dest_imp    = dest_binop "op -->"
-val dest_conj   = dest_binop "op &"
-val dest_disj   = dest_binop "op |"
-val dest_cons   = dest_binop "Cons"
-val dest_let    = Library.swap o dest_binop "Let";
-val dest_select = dest_binder "Hilbert_Choice.Eps"
-val dest_exists = dest_binder "Ex"
-val dest_forall = dest_binder "All"
+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_select = dest_binder @{const_name Eps}
+val dest_exists = dest_binder @{const_name Ex}
+val dest_forall = dest_binder @{const_name All}
 
 (* Query routines *)
 
@@ -151,8 +146,6 @@
 val is_conj   = can dest_conj
 val is_disj   = can dest_disj
 val is_pair   = can dest_pair
-val is_let    = can dest_let
-val is_cons   = can dest_cons
 
 
 (*---------------------------------------------------------------------------
@@ -197,7 +190,7 @@
   handle TYPE (msg, _, _) => raise ERR "mk_prop" msg
     | TERM (msg, _) => raise ERR "mk_prop" msg;
 
-fun drop_prop ctm = dest_monop "Trueprop" ctm handle U.ERR _ => ctm;
+fun drop_prop ctm = dest_monop @{const_name Trueprop} ctm handle U.ERR _ => ctm;
 
 
 end;
--- a/src/HOL/Tools/TFL/post.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/TFL/post.ML	Thu Aug 19 17:41:52 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("op =",_)$_$_) => r RS eq_reflection
+  |   _ $(Const(@{const_name "op ="},_)$_$_) => 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 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/TFL/rules.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -453,14 +453,14 @@
 (* Fragile: it's a cong if it is not "R y x ==> cut f R x y = f y" *)
 fun is_cong thm =
   case (Thm.prop_of thm)
-     of (Const("==>",_)$(Const("Trueprop",_)$ _) $
+     of (Const("==>",_)$(Const(@{const_name Trueprop},_)$ _) $
          (Const("==",_) $ (Const (@{const_name Recdef.cut},_) $ f $ R $ a $ x) $ _)) => false
       | _ => true;
 
 
 fun dest_equal(Const ("==",_) $
-               (Const ("Trueprop",_) $ lhs)
-               $ (Const ("Trueprop",_) $ rhs)) = {lhs=lhs, rhs=rhs}
+               (Const (@{const_name Trueprop},_) $ lhs)
+               $ (Const (@{const_name Trueprop},_) $ rhs)) = {lhs=lhs, rhs=rhs}
   | dest_equal(Const ("==",_) $ lhs $ rhs)  = {lhs=lhs, rhs=rhs}
   | dest_equal tm = S.dest_eq tm;
 
@@ -759,7 +759,7 @@
               val cntxt = rev(Simplifier.prems_of_ss ss)
               val dummy = print_thms "cntxt:" cntxt
               val thy = Thm.theory_of_thm thm
-              val Const("==>",_) $ (Const("Trueprop",_) $ A) $ _ = Thm.prop_of thm
+              val Const("==>",_) $ (Const(@{const_name Trueprop},_) $ A) $ _ = Thm.prop_of thm
               fun genl tm = let val vlist = subtract (op aconv) globals
                                            (OldTerm.add_term_frees(tm,[]))
                             in fold_rev Forall vlist tm
--- a/src/HOL/Tools/TFL/tfl.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/TFL/tfl.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -483,7 +483,7 @@
      val (case_rewrites,context_congs) = extraction_thms thy
      val tych = Thry.typecheck thy
      val WFREC_THM0 = R.ISPEC (tych functional) Thms.WFREC_COROLLARY
-     val Const("All",_) $ Abs(Rname,Rtype,_) = concl WFREC_THM0
+     val Const(@{const_name All},_) $ Abs(Rname,Rtype,_) = concl WFREC_THM0
      val R = Free (Name.variant (List.foldr OldTerm.add_term_names [] eqns) Rname,
                    Rtype)
      val WFREC_THM = R.ISPECL [tych R, tych g] WFREC_THM0
--- a/src/HOL/Tools/TFL/usyntax.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/TFL/usyntax.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -159,36 +159,36 @@
 
 
 fun mk_imp{ant,conseq} =
-   let val c = Const("op -->",HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
+   let val c = Const(@{const_name "op -->"},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
    in list_comb(c,[ant,conseq])
    end;
 
 fun mk_select (r as {Bvar,Body}) =
   let val ty = type_of Bvar
-      val c = Const("Hilbert_Choice.Eps",(ty --> HOLogic.boolT) --> ty)
+      val c = Const(@{const_name Eps},(ty --> HOLogic.boolT) --> ty)
   in list_comb(c,[mk_abs r])
   end;
 
 fun mk_forall (r as {Bvar,Body}) =
   let val ty = type_of Bvar
-      val c = Const("All",(ty --> HOLogic.boolT) --> HOLogic.boolT)
+      val c = Const(@{const_name All},(ty --> HOLogic.boolT) --> HOLogic.boolT)
   in list_comb(c,[mk_abs r])
   end;
 
 fun mk_exists (r as {Bvar,Body}) =
   let val ty = type_of Bvar
-      val c = Const("Ex",(ty --> HOLogic.boolT) --> HOLogic.boolT)
+      val c = Const(@{const_name Ex},(ty --> HOLogic.boolT) --> HOLogic.boolT)
   in list_comb(c,[mk_abs r])
   end;
 
 
 fun mk_conj{conj1,conj2} =
-   let val c = Const("op &",HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
+   let val c = Const(@{const_name "op &"},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
    in list_comb(c,[conj1,conj2])
    end;
 
 fun mk_disj{disj1,disj2} =
-   let val c = Const("op |",HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
+   let val c = Const(@{const_name "op |"},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
    in list_comb(c,[disj1,disj2])
    end;
 
@@ -244,25 +244,25 @@
      end
   | dest_abs _ _ =  raise USYN_ERR "dest_abs" "not an abstraction";
 
-fun dest_eq(Const("op =",_) $ M $ N) = {lhs=M, rhs=N}
+fun dest_eq(Const(@{const_name "op ="},_) $ M $ N) = {lhs=M, rhs=N}
   | dest_eq _ = raise USYN_ERR "dest_eq" "not an equality";
 
-fun dest_imp(Const("op -->",_) $ M $ N) = {ant=M, conseq=N}
+fun dest_imp(Const(@{const_name "op -->"},_) $ M $ N) = {ant=M, conseq=N}
   | dest_imp _ = raise USYN_ERR "dest_imp" "not an implication";
 
-fun dest_forall(Const("All",_) $ (a as Abs _)) = fst (dest_abs [] a)
+fun dest_forall(Const(@{const_name All},_) $ (a as Abs _)) = fst (dest_abs [] a)
   | dest_forall _ = raise USYN_ERR "dest_forall" "not a forall";
 
-fun dest_exists(Const("Ex",_) $ (a as Abs _)) = fst (dest_abs [] a)
+fun dest_exists(Const(@{const_name Ex},_) $ (a as Abs _)) = fst (dest_abs [] a)
   | dest_exists _ = raise USYN_ERR "dest_exists" "not an existential";
 
-fun dest_neg(Const("not",_) $ M) = M
+fun dest_neg(Const(@{const_name Not},_) $ M) = M
   | dest_neg _ = raise USYN_ERR "dest_neg" "not a negation";
 
-fun dest_conj(Const("op &",_) $ M $ N) = {conj1=M, conj2=N}
+fun dest_conj(Const(@{const_name "op &"},_) $ M $ N) = {conj1=M, conj2=N}
   | dest_conj _ = raise USYN_ERR "dest_conj" "not a conjunction";
 
-fun dest_disj(Const("op |",_) $ M $ N) = {disj1=M, disj2=N}
+fun dest_disj(Const(@{const_name "op |"},_) $ M $ N) = {disj1=M, disj2=N}
   | dest_disj _ = raise USYN_ERR "dest_disj" "not a disjunction";
 
 fun mk_pair{fst,snd} =
@@ -402,6 +402,6 @@
   | is_WFR _                 = false;
 
 fun ARB ty = mk_select{Bvar=Free("v",ty),
-                       Body=Const("True",HOLogic.boolT)};
+                       Body=Const(@{const_name True},HOLogic.boolT)};
 
 end;
--- a/src/HOL/Tools/choice_specification.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/choice_specification.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -24,7 +24,7 @@
     fun mk_definitional [] arg = arg
       | mk_definitional ((thname,cname,covld)::cos) (thy,thm) =
         case HOLogic.dest_Trueprop (concl_of thm) of
-            Const("Ex",_) $ P =>
+            Const(@{const_name Ex},_) $ P =>
             let
                 val ctype = domain_type (type_of P)
                 val cname_full = Sign.intern_const thy cname
@@ -51,7 +51,7 @@
                 end
               | process ((thname,cname,covld)::cos) (thy,tm) =
                 case tm of
-                    Const("Ex",_) $ P =>
+                    Const(@{const_name Ex},_) $ P =>
                     let
                         val ctype = domain_type (type_of P)
                         val cname_full = Sign.intern_const thy cname
--- a/src/HOL/Tools/cnf_funcs.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/cnf_funcs.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -93,19 +93,19 @@
 
 val cnftac_eq_imp        = @{lemma "[| P = Q; P |] ==> Q" by auto};
 
-fun is_atom (Const ("False", _))                                           = false
-  | is_atom (Const ("True", _))                                            = false
-  | is_atom (Const ("op &", _) $ _ $ _)                                    = false
-  | is_atom (Const ("op |", _) $ _ $ _)                                    = false
-  | is_atom (Const ("op -->", _) $ _ $ _)                                  = false
-  | is_atom (Const ("op =", Type ("fun", Type ("bool", []) :: _)) $ _ $ _) = false
-  | is_atom (Const ("Not", _) $ _)                                         = false
+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 Not}, _) $ _)                                         = false
   | is_atom _                                                              = true;
 
-fun is_literal (Const ("Not", _) $ x) = is_atom x
+fun is_literal (Const (@{const_name Not}, _) $ x) = is_atom x
   | is_literal x                      = is_atom x;
 
-fun is_clause (Const ("op |", _) $ x $ y) = is_clause x andalso is_clause y
+fun is_clause (Const (@{const_name "op |"}, _) $ x $ y) = is_clause x andalso is_clause y
   | is_clause x                           = is_literal x;
 
 (* ------------------------------------------------------------------------- *)
@@ -118,7 +118,7 @@
 fun clause_is_trivial c =
 	let
 		(* Term.term -> Term.term *)
-		fun dual (Const ("Not", _) $ x) = x
+		fun dual (Const (@{const_name Not}, _) $ x) = x
 		  | dual x                      = HOLogic.Not $ x
 		(* Term.term list -> bool *)
 		fun has_duals []      = false
@@ -184,28 +184,28 @@
 
 (* Theory.theory -> Term.term -> Thm.thm *)
 
-fun make_nnf_thm thy (Const ("op &", _) $ x $ y) =
+fun make_nnf_thm thy (Const (@{const_name "op &"}, _) $ 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 ("op |", _) $ x $ y) =
+  | make_nnf_thm thy (Const (@{const_name "op |"}, _) $ 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 ("op -->", _) $ x $ y) =
+  | make_nnf_thm thy (Const (@{const_name "op -->"}, _) $ 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 ("op =", Type ("fun", Type ("bool", []) :: _)) $ x $ y) =
+  | make_nnf_thm thy (Const (@{const_name "op ="}, Type ("fun", @{typ bool} :: _)) $ x $ y) =
 	let
 		val thm1 = make_nnf_thm thy x
 		val thm2 = make_nnf_thm thy (HOLogic.Not $ x)
@@ -214,32 +214,32 @@
 	in
 		make_nnf_iff OF [thm1, thm2, thm3, thm4]
 	end
-  | make_nnf_thm thy (Const ("Not", _) $ Const ("False", _)) =
+  | make_nnf_thm thy (Const (@{const_name Not}, _) $ Const (@{const_name False}, _)) =
 	make_nnf_not_false
-  | make_nnf_thm thy (Const ("Not", _) $ Const ("True", _)) =
+  | make_nnf_thm thy (Const (@{const_name Not}, _) $ Const (@{const_name True}, _)) =
 	make_nnf_not_true
-  | make_nnf_thm thy (Const ("Not", _) $ (Const ("op &", _) $ x $ y)) =
+  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name "op &"}, _) $ 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 ("Not", _) $ (Const ("op |", _) $ x $ y)) =
+  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name "op |"}, _) $ 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 ("Not", _) $ (Const ("op -->", _) $ x $ y)) =
+  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name "op -->"}, _) $ 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 ("Not", _) $ (Const ("op =", Type ("fun", Type ("bool", []) :: _)) $ x $ y)) =
+  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name "op ="}, Type ("fun", @{typ bool} :: _)) $ x $ y)) =
 	let
 		val thm1 = make_nnf_thm thy x
 		val thm2 = make_nnf_thm thy (HOLogic.Not $ x)
@@ -248,7 +248,7 @@
 	in
 		make_nnf_not_iff OF [thm1, thm2, thm3, thm4]
 	end
-  | make_nnf_thm thy (Const ("Not", _) $ (Const ("Not", _) $ x)) =
+  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name Not}, _) $ x)) =
 	let
 		val thm1 = make_nnf_thm thy x
 	in
@@ -276,7 +276,7 @@
 
 (* Theory.theory -> Term.term -> Thm.thm *)
 
-fun simp_True_False_thm thy (Const ("op &", _) $ x $ y) =
+fun simp_True_False_thm thy (Const (@{const_name "op &"}, _) $ 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 ("op |", _) $ x $ y) =
+  | simp_True_False_thm thy (Const (@{const_name "op |"}, _) $ 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 ("op &", _) $ x $ y) =
+	fun make_cnf_thm_from_nnf (Const (@{const_name "op &"}, _) $ 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 ("op |", _) $ x $ y) =
+	  | make_cnf_thm_from_nnf (Const (@{const_name "op |"}, _) $ x $ y) =
 		let
 			(* produces a theorem "(x' | y') = t'", where x', y', and t' are in CNF *)
-			fun make_cnf_disj_thm (Const ("op &", _) $ x1 $ x2) y' =
+			fun make_cnf_disj_thm (Const (@{const_name "op &"}, _) $ 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 ("op &", _) $ y1 $ y2) =
+			  | make_cnf_disj_thm x' (Const (@{const_name "op &"}, _) $ y1 $ y2) =
 				let
 					val thm1 = make_cnf_disj_thm x' y1
 					val thm2 = make_cnf_disj_thm x' y2
@@ -403,34 +403,34 @@
 	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 ("op &", _) $ x $ y) : thm =
+	fun make_cnfx_thm_from_nnf (Const (@{const_name "op &"}, _) $ 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 ("op |", _) $ x $ y) =
+	  | make_cnfx_thm_from_nnf (Const (@{const_name "op |"}, _) $ 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 ("op &", _) $ x1 $ x2) y' =
+			fun make_cnfx_disj_thm (Const (@{const_name "op &"}, _) $ 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 ("op &", _) $ y1 $ y2) =
+			  | make_cnfx_disj_thm x' (Const (@{const_name "op &"}, _) $ y1 $ y2) =
 				let
 					val thm1 = make_cnfx_disj_thm x' y1
 					val thm2 = make_cnfx_disj_thm x' y2
 				in
 					make_cnf_disj_conj_r OF [thm1, thm2]  (* (x' | (y1 & y2)) = ((x' | y1)' & (x' | y2)') *)
 				end
-			  | make_cnfx_disj_thm (Const ("Ex", _) $ x') y' =
+			  | make_cnfx_disj_thm (Const (@{const_name Ex}, _) $ x') y' =
 				let
 					val thm1 = inst_thm thy [x', y'] make_cnfx_disj_ex_l   (* ((Ex x') | y') = (Ex (x' | y')) *)
 					val var  = new_free ()
@@ -441,7 +441,7 @@
 				in
 					iff_trans OF [thm1, thm5]  (* ((Ex x') | y') = (Ex v. body') *)
 				end
-			  | make_cnfx_disj_thm x' (Const ("Ex", _) $ y') =
+			  | make_cnfx_disj_thm x' (Const (@{const_name Ex}, _) $ y') =
 				let
 					val thm1 = inst_thm thy [x', y'] make_cnfx_disj_ex_r   (* (x' | (Ex y')) = (Ex (x' | y')) *)
 					val var  = new_free ()
--- a/src/HOL/Tools/groebner.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/groebner.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -395,7 +395,7 @@
 
 fun refute_disj rfn tm =
  case term_of tm of
-  Const("op |",_)$l$r =>
+  Const(@{const_name "op |"},_)$l$r =>
    compose_single(refute_disj rfn (dest_arg tm),2,compose_single(refute_disj rfn (dest_arg1 tm),2,disjE))
   | _ => rfn tm ;
 
@@ -405,11 +405,11 @@
 val mk_comb = capply;
 fun is_neg t =
     case term_of t of
-      (Const("Not",_)$p) => true
+      (Const(@{const_name Not},_)$p) => true
     | _  => false;
 fun is_eq t =
  case term_of t of
- (Const("op =",_)$_$_) => true
+ (Const(@{const_name "op ="},_)$_$_) => true
 | _  => false;
 
 fun end_itlist f l =
@@ -430,14 +430,14 @@
 val strip_exists =
  let fun h (acc, t) =
       case (term_of t) of
-       Const("Ex",_)$Abs(x,T,p) => h (dest_abs NONE (dest_arg t) |>> (fn v => v::acc))
+       Const(@{const_name Ex},_)$Abs(x,T,p) => h (dest_abs NONE (dest_arg t) |>> (fn v => v::acc))
      | _ => (acc,t)
  in fn t => h ([],t)
  end;
 
 fun is_forall t =
  case term_of t of
-  (Const("All",_)$Abs(_,_,_)) => true
+  (Const(@{const_name All},_)$Abs(_,_,_)) => true
 | _ => false;
 
 val mk_object_eq = fn th => th COMP meta_eq_to_obj_eq;
@@ -522,7 +522,7 @@
  fun mk_ex v t = Thm.capply (ext (ctyp_of_term v)) (Thm.cabs v t)
 
 fun choose v th th' = case concl_of th of 
-  @{term Trueprop} $ (Const("Ex",_)$_) => 
+  @{term Trueprop} $ (Const(@{const_name Ex},_)$_) => 
    let
     val p = (funpow 2 Thm.dest_arg o cprop_of) th
     val T = (hd o Thm.dest_ctyp o ctyp_of_term) p
@@ -923,15 +923,15 @@
 
 fun find_term bounds tm =
   (case term_of tm of
-    Const ("op =", T) $ _ $ _ =>
+    Const (@{const_name "op ="}, T) $ _ $ _ =>
       if domain_type T = HOLogic.boolT then find_args bounds tm
       else dest_arg tm
-  | Const ("Not", _) $ _ => find_term bounds (dest_arg tm)
-  | Const ("All", _) $ _ => find_body bounds (dest_arg tm)
-  | Const ("Ex", _) $ _ => find_body bounds (dest_arg tm)
-  | Const ("op &", _) $ _ $ _ => find_args bounds tm
-  | Const ("op |", _) $ _ $ _ => find_args bounds tm
-  | Const ("op -->", _) $ _ $ _ => find_args bounds 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
   | @{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("op =",_)$_$_ => Thm.dest_arg1 t
+  Const(@{const_name "op ="},_)$_$_ => 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 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/hologic.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -143,15 +143,15 @@
 
 (* bool and set *)
 
-val boolN = "bool";
+val boolN = "HOL.bool";
 val boolT = Type (boolN, []);
 
-val true_const =  Const ("True", boolT);
-val false_const = Const ("False", boolT);
+val true_const =  Const ("HOL.True", boolT);
+val false_const = Const ("HOL.False", boolT);
 
 fun mk_setT T = T --> boolT;
 
-fun dest_setT (Type ("fun", [T, Type ("bool", [])])) = T
+fun dest_setT (Type ("fun", [T, Type ("HOL.bool", [])])) = T
   | dest_setT T = raise TYPE ("dest_setT: set type expected", [T], []);
 
 fun mk_set T ts =
@@ -181,11 +181,11 @@
 
 (* logic *)
 
-val Trueprop = Const ("Trueprop", boolT --> propT);
+val Trueprop = Const ("HOL.Trueprop", boolT --> propT);
 
 fun mk_Trueprop P = Trueprop $ P;
 
-fun dest_Trueprop (Const ("Trueprop", _) $ P) = P
+fun dest_Trueprop (Const ("HOL.Trueprop", _) $ P) = P
   | dest_Trueprop t = raise TERM ("dest_Trueprop", [t]);
 
 fun conj_intr thP thQ =
@@ -230,23 +230,23 @@
 
 fun disjuncts t = disjuncts_aux t [];
 
-fun dest_imp (Const("op -->",_) $ A $ B) = (A, B)
+fun dest_imp (Const ("op -->", _) $ A $ B) = (A, B)
   | dest_imp  t = raise TERM ("dest_imp", [t]);
 
-fun dest_not (Const ("Not", _) $ t) = 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 ("op =", T --> T --> boolT);
 fun mk_eq (t, u) = eq_const (fastype_of t) $ t $ u;
 
 fun dest_eq (Const ("op =", _) $ lhs $ rhs) = (lhs, rhs)
   | dest_eq t = raise TERM ("dest_eq", [t])
 
-fun all_const T = Const ("All", [T --> boolT] ---> boolT);
+fun all_const T = Const ("HOL.All", [T --> boolT] ---> boolT);
 fun mk_all (x, T, P) = all_const T $ absfree (x, T, P);
 fun list_all (xs, t) = fold_rev (fn (x, T) => fn P => all_const T $ Abs (x, T, P)) xs t;
 
-fun exists_const T = Const ("Ex", [T --> boolT] ---> boolT);
+fun exists_const T = Const ("HOL.Ex", [T --> boolT] ---> boolT);
 fun mk_exists (x, T, P) = exists_const T $ absfree (x, T, P);
 
 fun choice_const T = Const("Hilbert_Choice.Eps", (T --> boolT) --> T);
@@ -657,7 +657,9 @@
             $ t $ t', U)
   in fold_rev mk_clause clauses (t, U) |> fst end;
 
-val code_numeralT = Type ("Code_Numeral.code_numeral", []);
+
+(* random seeds *)
+
 val random_seedT = mk_prodT (code_numeralT, code_numeralT);
 
 fun mk_random T t = Const ("Quickcheck.random_class.random", code_numeralT
--- a/src/HOL/Tools/lin_arith.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/lin_arith.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -46,12 +46,12 @@
 val mk_Trueprop = HOLogic.mk_Trueprop;
 
 fun atomize thm = case Thm.prop_of thm of
-    Const ("Trueprop", _) $ (Const (@{const_name "op &"}, _) $ _ $ _) =>
+    Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op &"}, _) $ _ $ _) =>
     atomize (thm RS conjunct1) @ atomize (thm RS conjunct2)
   | _ => [thm];
 
-fun neg_prop ((TP as Const("Trueprop", _)) $ (Const (@{const_name Not}, _) $ t)) = TP $ t
-  | neg_prop ((TP as Const("Trueprop", _)) $ t) = TP $ (HOLogic.Not $t)
+fun neg_prop ((TP as Const(@{const_name Trueprop}, _)) $ (Const (@{const_name Not}, _) $ t)) = TP $ t
+  | neg_prop ((TP as Const(@{const_name Trueprop}, _)) $ t) = TP $ (HOLogic.Not $t)
   | neg_prop t = raise TERM ("neg_prop", [t]);
 
 fun is_False thm =
@@ -258,10 +258,10 @@
   | negate NONE                        = NONE;
 
 fun decomp_negation data
-  ((Const ("Trueprop", _)) $ (Const (rel, T) $ lhs $ rhs)) : decomp option =
+  ((Const (@{const_name Trueprop}, _)) $ (Const (rel, T) $ lhs $ rhs)) : decomp option =
       decomp_typecheck data (T, (rel, lhs, rhs))
-  | decomp_negation data ((Const ("Trueprop", _)) $
-  (Const ("Not", _) $ (Const (rel, T) $ lhs $ rhs))) =
+  | decomp_negation data ((Const (@{const_name Trueprop}, _)) $
+  (Const (@{const_name Not}, _) $ (Const (rel, T) $ lhs $ rhs))) =
       negate (decomp_typecheck data (T, (rel, lhs, rhs)))
   | decomp_negation data _ =
       NONE;
@@ -273,7 +273,7 @@
   in decomp_negation (thy, discrete, inj_consts) end;
 
 fun domain_is_nat (_ $ (Const (_, T) $ _ $ _))                      = nT T
-  | domain_is_nat (_ $ (Const ("Not", _) $ (Const (_, T) $ _ $ _))) = nT T
+  | domain_is_nat (_ $ (Const (@{const_name Not}, _) $ (Const (_, T) $ _ $ _))) = nT T
   | domain_is_nat _                                                 = false;
 
 
@@ -428,7 +428,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 ("op =", split_type --> split_type --> HOLogic.boolT) $ t1' $
+        val t1_eq_t2_plus_d = Const (@{const_name "op ="}, 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 +448,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 ("op =", HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ t1' $
+        val t1_eq_nat_n = Const (@{const_name "op ="}, 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 +472,13 @@
                                         (map (incr_boundvars 2) rev_terms)
         val t1'                     = incr_boundvars 2 t1
         val t2'                     = incr_boundvars 2 t2
-        val t2_eq_zero              = Const ("op =",
+        val t2_eq_zero              = Const (@{const_name "op ="},
                                         split_type --> split_type --> HOLogic.boolT) $ t2 $ zero
-        val t2_neq_zero             = HOLogic.mk_not (Const ("op =",
+        val t2_neq_zero             = HOLogic.mk_not (Const (@{const_name "op ="},
                                         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 ("op =", split_type --> split_type --> HOLogic.boolT) $ t1' $
+        val t1_eq_t2_times_i_plus_j = Const (@{const_name "op ="}, 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 +504,13 @@
                                         (map (incr_boundvars 2) rev_terms)
         val t1'                     = incr_boundvars 2 t1
         val t2'                     = incr_boundvars 2 t2
-        val t2_eq_zero              = Const ("op =",
+        val t2_eq_zero              = Const (@{const_name "op ="},
                                         split_type --> split_type --> HOLogic.boolT) $ t2 $ zero
-        val t2_neq_zero             = HOLogic.mk_not (Const ("op =",
+        val t2_neq_zero             = HOLogic.mk_not (Const (@{const_name "op ="},
                                         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 ("op =", split_type --> split_type --> HOLogic.boolT) $ t1' $
+        val t1_eq_t2_times_i_plus_j = Const (@{const_name "op ="}, 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 +542,7 @@
                                         (map (incr_boundvars 2) rev_terms)
         val t1'                     = incr_boundvars 2 t1
         val t2'                     = incr_boundvars 2 t2
-        val t2_eq_zero              = Const ("op =",
+        val t2_eq_zero              = Const (@{const_name "op ="},
                                         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 +556,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 ("op =", split_type --> split_type --> HOLogic.boolT) $ t1' $
+        val t1_eq_t2_times_i_plus_j = Const (@{const_name "op ="}, 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 +596,7 @@
                                         (map (incr_boundvars 2) rev_terms)
         val t1'                     = incr_boundvars 2 t1
         val t2'                     = incr_boundvars 2 t2
-        val t2_eq_zero              = Const ("op =",
+        val t2_eq_zero              = Const (@{const_name "op ="},
                                         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 +610,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 ("op =", split_type --> split_type --> HOLogic.boolT) $ t1' $
+        val t1_eq_t2_times_i_plus_j = Const (@{const_name "op ="}, 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)
@@ -659,7 +659,7 @@
 
 fun negated_term_occurs_positively (terms : term list) : bool =
   List.exists
-    (fn (Trueprop $ (Const ("Not", _) $ t)) =>
+    (fn (Trueprop $ (Const (@{const_name Not}, _) $ t)) =>
       member Pattern.aeconv terms (Trueprop $ t)
       | _ => false)
     terms;
--- a/src/HOL/Tools/meson.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/meson.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -149,21 +149,21 @@
 (*Are any of the logical connectives in "bs" present in the term?*)
 fun has_conns bs =
   let fun has (Const(a,_)) = false
-        | has (Const("Trueprop",_) $ p) = has p
-        | has (Const("Not",_) $ p) = has p
-        | has (Const("op |",_) $ p $ q) = member (op =) bs "op |" orelse has p orelse has q
-        | has (Const("op &",_) $ p $ q) = member (op =) bs "op &" orelse has p orelse has q
-        | has (Const("All",_) $ Abs(_,_,p)) = member (op =) bs "All" orelse has p
-        | has (Const("Ex",_) $ Abs(_,_,p)) = member (op =) bs "Ex" orelse has p
+        | 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 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
   in  has  end;
 
 
 (**** Clause handling ****)
 
-fun literals (Const("Trueprop",_) $ P) = literals P
-  | literals (Const("op |",_) $ P $ Q) = literals P @ literals Q
-  | literals (Const("Not",_) $ P) = [(false,P)]
+fun literals (Const(@{const_name Trueprop},_) $ P) = literals P
+  | literals (Const(@{const_name "op |"},_) $ P $ Q) = literals P @ literals Q
+  | literals (Const(@{const_name Not},_) $ P) = [(false,P)]
   | literals P = [(true,P)];
 
 (*number of literals in a term*)
@@ -172,16 +172,16 @@
 
 (*** Tautology Checking ***)
 
-fun signed_lits_aux (Const ("op |", _) $ P $ Q) (poslits, neglits) =
+fun signed_lits_aux (Const (@{const_name "op |"}, _) $ P $ Q) (poslits, neglits) =
       signed_lits_aux Q (signed_lits_aux P (poslits, neglits))
-  | signed_lits_aux (Const("Not",_) $ P) (poslits, neglits) = (poslits, P::neglits)
+  | signed_lits_aux (Const(@{const_name Not},_) $ P) (poslits, neglits) = (poslits, P::neglits)
   | signed_lits_aux P (poslits, neglits) = (P::poslits, neglits);
 
 fun signed_lits th = signed_lits_aux (HOLogic.dest_Trueprop (concl_of th)) ([],[]);
 
 (*Literals like X=X are tautologous*)
-fun taut_poslit (Const("op =",_) $ t $ u) = t aconv u
-  | taut_poslit (Const("True",_)) = true
+fun taut_poslit (Const(@{const_name "op ="},_) $ t $ u) = t aconv u
+  | taut_poslit (Const(@{const_name True},_)) = true
   | taut_poslit _ = false;
 
 fun is_taut th =
@@ -208,18 +208,18 @@
 fun refl_clause_aux 0 th = th
   | refl_clause_aux n th =
        case HOLogic.dest_Trueprop (concl_of th) of
-          (Const ("op |", _) $ (Const ("op |", _) $ _ $ _) $ _) =>
+          (Const (@{const_name "op |"}, _) $ (Const (@{const_name "op |"}, _) $ _ $ _) $ _) =>
             refl_clause_aux n (th RS disj_assoc)    (*isolate an atom as first disjunct*)
-        | (Const ("op |", _) $ (Const("Not",_) $ (Const("op =",_) $ t $ u)) $ _) =>
+        | (Const (@{const_name "op |"}, _) $ (Const(@{const_name Not},_) $ (Const(@{const_name "op ="},_) $ 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 ("op |", _) $ _ $ _) => refl_clause_aux n (th RS disj_comm)
+        | (Const (@{const_name "op |"}, _) $ _ $ _) => refl_clause_aux n (th RS disj_comm)
         | _ => (*not a disjunction*) th;
 
-fun notequal_lits_count (Const ("op |", _) $ P $ Q) =
+fun notequal_lits_count (Const (@{const_name "op |"}, _) $ P $ Q) =
       notequal_lits_count P + notequal_lits_count Q
-  | notequal_lits_count (Const("Not",_) $ (Const("op =",_) $ _ $ _)) = 1
+  | notequal_lits_count (Const(@{const_name Not},_) $ (Const(@{const_name "op ="},_) $ _ $ _)) = 1
   | notequal_lits_count _ = 0;
 
 (*Simplify a clause by applying reflexivity to its negated equality literals*)
@@ -266,26 +266,26 @@
   fun prod x y = if x < max_cl andalso y < max_cl then x*y else max_cl;
   
   (*Estimate the number of clauses in order to detect infeasible theorems*)
-  fun signed_nclauses b (Const("Trueprop",_) $ t) = signed_nclauses b t
-    | signed_nclauses b (Const("Not",_) $ t) = signed_nclauses (not b) t
-    | signed_nclauses b (Const("op &",_) $ t $ u) =
+  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) =
         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("op |",_) $ t $ u) =
+    | signed_nclauses b (Const(@{const_name "op |"},_) $ 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("op -->",_) $ t $ u) =
+    | signed_nclauses b (Const(@{const_name "op -->"},_) $ 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("op =", Type ("fun", [T, _])) $ t $ u) =
+    | signed_nclauses b (Const(@{const_name "op ="}, 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))
                  else sum (prod (signed_nclauses b t) (signed_nclauses b u))
                           (prod (signed_nclauses (not b) t) (signed_nclauses (not b) u))
         else 1
-    | signed_nclauses b (Const("Ex", _) $ Abs (_,_,t)) = signed_nclauses b t
-    | signed_nclauses b (Const("All",_) $ Abs (_,_,t)) = signed_nclauses b t
+    | signed_nclauses b (Const(@{const_name Ex}, _) $ Abs (_,_,t)) = signed_nclauses b t
+    | signed_nclauses b (Const(@{const_name All},_) $ Abs (_,_,t)) = signed_nclauses b t
     | signed_nclauses _ _ = 1; (* literal *)
  in 
   signed_nclauses true t >= max_cl
@@ -296,7 +296,7 @@
 local  
   val spec_var = Thm.dest_arg (Thm.dest_arg (#2 (Thm.dest_implies (Thm.cprop_of spec))));
   val spec_varT = #T (Thm.rep_cterm spec_var);
-  fun name_of (Const ("All", _) $ Abs(x,_,_)) = x | name_of _ = Name.uu;
+  fun name_of (Const (@{const_name All}, _) $ Abs(x,_,_)) = x | name_of _ = Name.uu;
 in  
   fun freeze_spec th ctxt =
     let
@@ -314,8 +314,7 @@
 
 (*Any need to extend this list with
   "HOL.type_class","HOL.eq_class","Pure.term"?*)
-val has_meta_conn =
-    exists_Const (member (op =) ["==", "==>", "=simp=>", "all", "prop"] o #1);
+val has_meta_conn = exists_Const (member (op =) ["==", "==>", "=simp=>", "all", "prop"] o #1);
 
 fun apply_skolem_theorem (th, rls) =
   let
@@ -331,18 +330,18 @@
   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 ["All","Ex","op &"] (prop_of th))
+        else if not (has_conns [@{const_name All}, @{const_name Ex}, @{const_name "op &"}] (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 ("op &", _) => (*conjunction*)
+            Const (@{const_name "op &"}, _) => (*conjunction*)
                 cnf_aux (th RS conjunct1, cnf_aux (th RS conjunct2, ths))
-          | Const ("All", _) => (*universal quantifier*)
+          | Const (@{const_name All}, _) => (*universal quantifier*)
                 let val (th',ctxt') = freeze_spec th (!ctxtr)
                 in  ctxtr := ctxt'; cnf_aux (th', ths) end
-          | Const ("Ex", _) =>
+          | Const (@{const_name Ex}, _) =>
               (*existential quantifier: Insert Skolem functions*)
               cnf_aux (apply_skolem_theorem (th, skolem_ths), ths)
-          | Const ("op |", _) =>
+          | Const (@{const_name "op |"}, _) =>
               (*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,8 +364,8 @@
 
 (**** Generation of contrapositives ****)
 
-fun is_left (Const ("Trueprop", _) $
-               (Const ("op |", _) $ (Const ("op |", _) $ _ $ _) $ _)) = true
+fun is_left (Const (@{const_name Trueprop}, _) $
+               (Const (@{const_name "op |"}, _) $ (Const (@{const_name "op |"}, _) $ _ $ _) $ _)) = true
   | is_left _ = false;
 
 (*Associate disjuctions to right -- make leftmost disjunct a LITERAL*)
@@ -401,8 +400,9 @@
 (*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 =)
-    ["Trueprop", "op &", "op |", "op -->", "Not",
-     "All", "Ex", @{const_name Ball}, @{const_name Bex}];
+    [@{const_name Trueprop}, @{const_name "op &"}, @{const_name "op |"},
+     @{const_name "op -->"}, @{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
   of any argument contains bool.*)
@@ -420,7 +420,7 @@
 (*Returns false if any Vars in the theorem mention type bool.
   Also rejects functions whose arguments are Booleans or other functions.*)
 fun is_fol_term thy t =
-    Term.is_first_order ["all","All","Ex"] t andalso
+    Term.is_first_order ["all", @{const_name All}, @{const_name Ex}] t andalso
     not (exists_subterm (fn Var (_, T) => has_bool T orelse has_fun T
                            | _ => false) t orelse
          has_bool_arg_const t orelse
@@ -429,8 +429,8 @@
 
 fun rigid t = not (is_Var (head_of t));
 
-fun ok4horn (Const ("Trueprop",_) $ (Const ("op |", _) $ t $ _)) = rigid t
-  | ok4horn (Const ("Trueprop",_) $ t) = rigid t
+fun ok4horn (Const (@{const_name Trueprop},_) $ (Const (@{const_name "op |"}, _) $ t $ _)) = rigid t
+  | ok4horn (Const (@{const_name Trueprop},_) $ t) = rigid t
   | ok4horn _ = false;
 
 (*Create a meta-level Horn clause*)
@@ -464,7 +464,7 @@
 
 (***** MESON PROOF PROCEDURE *****)
 
-fun rhyps (Const("==>",_) $ (Const("Trueprop",_) $ A) $ phi,
+fun rhyps (Const("==>",_) $ (Const(@{const_name Trueprop},_) $ A) $ phi,
            As) = rhyps(phi, A::As)
   | rhyps (_, As) = As;
 
@@ -509,8 +509,8 @@
 val nnf_rls = [imp_to_disjD, iff_to_disjD, not_conjD, not_disjD,
                not_impD, not_iffD, not_allD, not_exD, not_notD];
 
-fun ok4nnf (Const ("Trueprop",_) $ (Const ("Not", _) $ t)) = rigid t
-  | ok4nnf (Const ("Trueprop",_) $ t) = rigid t
+fun ok4nnf (Const (@{const_name Trueprop},_) $ (Const (@{const_name Not}, _) $ t)) = rigid t
+  | ok4nnf (Const (@{const_name Trueprop},_) $ t) = rigid t
   | ok4nnf _ = false;
 
 fun make_nnf1 ctxt th =
@@ -546,7 +546,7 @@
 (*Pull existential quantifiers to front. This accomplishes Skolemization for
   clauses that arise from a subgoal.*)
 fun skolemize1 ctxt th =
-  if not (has_conns ["Ex"] (prop_of th)) then th
+  if not (has_conns [@{const_name Ex}] (prop_of th)) then th
   else (skolemize1 ctxt (tryres(th, [choice, conj_exD1, conj_exD2,
                               disj_exD, disj_exD1, disj_exD2])))
     handle THM ("tryres", _, _) =>
--- a/src/HOL/Tools/numeral_simprocs.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/numeral_simprocs.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -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("op =",_)$(Const(@{const_name Rings.divide},_)$_$_)$_ =>
+  | Const(@{const_name "op ="},_)$(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("op =",_)$_$(Const(@{const_name Rings.divide},_)$_$_) =>
+  | Const(@{const_name "op ="},_)$_$(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 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/prop_logic.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -391,20 +391,20 @@
 				next_idx_is_valid := true;
 				Unsynchronized.inc next_idx
 			)
-		fun aux (Const ("True", _))         table =
+		fun aux (Const (@{const_name True}, _))         table =
 			(True, table)
-		  | aux (Const ("False", _))        table =
+		  | aux (Const (@{const_name False}, _))        table =
 			(False, table)
-		  | aux (Const ("Not", _) $ x)      table =
+		  | aux (Const (@{const_name Not}, _) $ x)      table =
 			apfst Not (aux x table)
-		  | aux (Const ("op |", _) $ x $ y) table =
+		  | aux (Const (@{const_name "op |"}, _) $ x $ y) table =
 			let
 				val (fm1, table1) = aux x table
 				val (fm2, table2) = aux y table1
 			in
 				(Or (fm1, fm2), table2)
 			end
-		  | aux (Const ("op &", _) $ x $ y) table =
+		  | aux (Const (@{const_name "op &"}, _) $ 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 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/quickcheck_generators.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -10,10 +10,8 @@
   val random_fun: typ -> typ -> ('a -> 'a -> bool) -> ('a -> term)
     -> (seed -> ('b * (unit -> term)) * seed) -> (seed -> seed * seed)
     -> seed -> (('a -> 'b) * (unit -> term)) * seed
-  val random_aux_specification: string -> string -> term list -> local_theory -> local_theory
-  val mk_random_aux_eqs: theory -> Datatype.descr -> (string * sort) list
-    -> string list -> string list * string list -> typ list * typ list
-    -> term list * (term * term) list
+  val perhaps_constrain: theory -> (typ * sort) list -> (string * sort) list
+    -> (string * sort -> string * sort) option
   val ensure_random_datatype: Datatype.config -> string list -> theory -> theory
   val compile_generator_expr:
     theory -> bool -> term -> int -> term list option * (bool list * bool)
@@ -219,11 +217,11 @@
         val is_rec = exists is_some ks;
         val k = fold (fn NONE => I | SOME k => Integer.max k) ks 0;
         val vs = Name.names Name.context "x" (map snd simple_tTs);
-        val vs' = (map o apsnd) termifyT vs;
         val tc = HOLogic.mk_return T @{typ Random.seed}
           (HOLogic.mk_valtermify_app c vs simpleT);
-        val t = HOLogic.mk_ST (map (fn (t, _) => (t, @{typ Random.seed})) tTs ~~ map SOME vs')
-          tc @{typ Random.seed} (SOME T, @{typ Random.seed});
+        val t = HOLogic.mk_ST
+          (map2 (fn (t, _) => fn (v, T') => ((t, @{typ Random.seed}), SOME ((v, termifyT T')))) tTs vs)
+            tc @{typ Random.seed} (SOME T, @{typ Random.seed});
         val tk = if is_rec
           then if k = 0 then size
             else @{term "Quickcheck.beyond :: code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"}
@@ -245,7 +243,7 @@
     val auxs_rhss = map mk_select gen_exprss;
   in (random_auxs, auxs_lhss ~~ auxs_rhss) end;
 
-fun mk_random_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy =
+fun instantiate_random_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy =
   let
     val _ = Datatype_Aux.message config "Creating quickcheck generators ...";
     val mk_prop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq;
@@ -272,11 +270,11 @@
 
 fun perhaps_constrain thy insts raw_vs =
   let
-    fun meet_random (T, sort) = Sorts.meet_sort (Sign.classes_of thy) 
+    fun meet (T, sort) = Sorts.meet_sort (Sign.classes_of thy) 
       (Logic.varifyT_global T, sort);
     val vtab = Vartab.empty
       |> fold (fn (v, sort) => Vartab.update ((v, 0), sort)) raw_vs
-      |> fold meet_random insts;
+      |> fold meet insts;
   in SOME (fn (v, _) => (v, (the o Vartab.lookup vtab) (v, 0)))
   end handle Sorts.CLASS_ERROR _ => NONE;
 
@@ -297,7 +295,7 @@
       can (Sorts.mg_domain algebra tyco) @{sort random}) tycos;
   in if has_inst then thy
     else case perhaps_constrain thy (random_insts @ term_of_insts) typerep_vs
-     of SOME constrain => mk_random_datatype config descr
+     of SOME constrain => instantiate_random_datatype config descr
           (map constrain typerep_vs) tycos prfx (names, auxnames)
             ((pairself o map o map_atyps) (fn TFree v => TFree (constrain v)) raw_TUs) thy
       | NONE => thy
@@ -344,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("op -->",_) $ A $ B) = apfst (cons A) (strip_imp B)
+    fun strip_imp (Const(@{const_name "op -->"},_) $ 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)
@@ -353,11 +351,11 @@
       @{term "Pair :: term list option * (bool list * bool) => Random.seed => (term list option * (bool list * bool)) * Random.seed"};
     fun mk_assms_report i =
       HOLogic.mk_prod (@{term "None :: term list option"},
-        HOLogic.mk_prod (HOLogic.mk_list @{typ "bool"}
-          (replicate i @{term "True"} @ replicate (length assms - i) @{term "False"}),
-        @{term "False"}))
+        HOLogic.mk_prod (HOLogic.mk_list HOLogic.boolT
+          (replicate i @{term True} @ replicate (length assms - i) @{term False}),
+        @{term False}))
     fun mk_concl_report b =
-      HOLogic.mk_prod (HOLogic.mk_list @{typ "bool"} (replicate (length assms) @{term "True"}),
+      HOLogic.mk_prod (HOLogic.mk_list HOLogic.boolT (replicate (length assms) @{term True}),
         if b then @{term True} else @{term False})
     val If =
       @{term "If :: bool => term list option * (bool list * bool) => term list option * (bool list * bool) => term list option * (bool list * bool)"}
@@ -385,22 +383,19 @@
 fun compile_generator_expr thy report t =
   let
     val Ts = (map snd o fst o strip_abs) t;
-  in
-    if report then
-      let
-        val t' = mk_reporting_generator_expr thy t Ts;
-        val compile = Code_Eval.eval (SOME target) ("Quickcheck_Generators.eval_report_ref", eval_report_ref)
-          (fn proc => fn g => fn s => g s #>> ((apfst o Option.map o map) proc)) thy t' [];
-      in
-        compile #> Random_Engine.run
-      end
-    else
-      let
-        val t' = mk_generator_expr thy t Ts;
-        val compile = Code_Eval.eval (SOME target) ("Quickcheck_Generators.eval_ref", eval_ref)
-          (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) thy t' [];
-        val dummy_report = ([], false)
-      in fn s => ((compile #> Random_Engine.run) s, dummy_report) end
+  in if report then
+    let
+      val t' = mk_reporting_generator_expr thy t Ts;
+      val compile = Code_Eval.eval (SOME target) ("Quickcheck_Generators.eval_report_ref", eval_report_ref)
+        (fn proc => fn g => fn s => g s #>> (apfst o Option.map o map) proc) thy t' [];
+    in compile #> Random_Engine.run end
+  else
+    let
+      val t' = mk_generator_expr thy t Ts;
+      val compile = Code_Eval.eval (SOME target) ("Quickcheck_Generators.eval_ref", eval_ref)
+        (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) thy t' [];
+      val dummy_report = ([], false)
+    in compile #> Random_Engine.run #> rpair dummy_report end
   end;
 
 
--- a/src/HOL/Tools/quickcheck_record.ML	Thu Aug 19 17:40:44 2010 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,60 +0,0 @@
-(*  Title:      HOL/Tools/quickcheck_record.ML
-    Author:     Florian Haftmann, TU Muenchen
-
-Quickcheck generators for records.
-*)
-
-signature QUICKCHECK_RECORD =
-sig
-  val ensure_random_typecopy: string -> theory -> theory
-  val setup: theory -> theory
-end;
-
-structure Quickcheck_Record : QUICKCHECK_RECORD =
-struct
-
-fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit => term"})
-val size = @{term "i::code_numeral"};
-
-fun mk_random_typecopy tyco vs constr T' thy =
-  let
-    val mk_const = curry (Sign.mk_const thy);
-    val Ts = map TFree vs;  
-    val T = Type (tyco, Ts);
-    val Tm = termifyT T;
-    val Tm' = termifyT T';
-    val v = "x";
-    val t_v = Free (v, Tm');
-    val t_constr = Const (constr, T' --> T);
-    val lhs = HOLogic.mk_random T size;
-    val rhs = HOLogic.mk_ST [((HOLogic.mk_random T' size, @{typ Random.seed}), SOME (v, Tm'))]
-      (HOLogic.mk_return Tm @{typ Random.seed}
-      (mk_const "Code_Evaluation.valapp" [T', T]
-        $ HOLogic.mk_prod (t_constr, Abs ("u", @{typ unit}, HOLogic.reflect_term t_constr)) $ t_v))
-      @{typ Random.seed} (SOME Tm, @{typ Random.seed});
-    val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
-  in   
-    thy
-    |> Class.instantiation ([tyco], vs, @{sort random})
-    |> `(fn lthy => Syntax.check_term lthy eq)
-    |-> (fn eq => Specification.definition (NONE, (apfst Binding.conceal Attrib.empty_binding, eq)))
-    |> snd
-    |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
-  end;
-
-fun ensure_random_typecopy tyco thy =
-  let
-    val SOME { vs = raw_vs, constr, typ = raw_T, ... } =
-      Typecopy.get_info thy tyco;
-    val constrain = curry (Sorts.inter_sort (Sign.classes_of thy));
-    val T = map_atyps (fn TFree (v, sort) =>
-      TFree (v, constrain sort @{sort random})) raw_T;
-    val vs' = Term.add_tfreesT T [];
-    val vs = map (fn (v, sort) =>
-      (v, the_default (constrain sort @{sort typerep}) (AList.lookup (op =) vs' v))) raw_vs;
-    val can_inst = Sign.of_sort thy (T, @{sort random});
-  in if can_inst then mk_random_typecopy tyco vs constr T thy else thy end;
-
-val setup = Typecopy.interpretation ensure_random_typecopy;
-
-end;
--- a/src/HOL/Tools/record.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/record.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -93,21 +93,29 @@
   fun merge data = Symtab.merge Thm.eq_thm_prop data;   (* FIXME handle Symtab.DUP ?? *)
 );
 
-fun do_typedef b repT alphas thy =
+fun get_typedef_info tyco vs (({ rep_type, Abs_name, Rep_name, ...},
+    { Abs_inject, Rep_inject, Abs_inverse, Rep_inverse, ... }) : Typedef.info) thy =
   let
-    val (b_constr, b_proj) = (Binding.prefix_name "make_" b, Binding.prefix_name "dest_" b);
-    fun get_thms thy tyco =
-      let
-        val SOME { vs, constr, typ = repT, proj_inject, proj_constr, ... } =
-          Typecopy.get_info thy tyco;
-        val absT = Type (tyco, map TFree vs);
-      in
-        ((proj_inject, proj_constr), Const (constr, repT --> absT), absT)
-      end;
+    val exists_thm =
+      UNIV_I
+      |> Drule.instantiate' [SOME (ctyp_of thy (Logic.varifyT_global rep_type))] [];
+    val proj_constr = Abs_inverse OF [exists_thm];
+    val absT = Type (tyco, map TFree vs);
   in
     thy
-    |> Typecopy.typecopy (b, alphas) repT b_constr b_proj
-    |-> (fn (tyco, _) => `(fn thy => get_thms thy tyco))
+    |> pair (tyco, ((Rep_inject, proj_constr), Const (Abs_name, rep_type --> absT), absT))
+  end
+
+fun do_typedef raw_tyco repT raw_vs thy =
+  let
+    val ctxt = ProofContext.init_global thy |> Variable.declare_typ repT;
+    val vs = map (ProofContext.check_tfree ctxt) raw_vs;
+    val tac = Tactic.rtac UNIV_witness 1;
+  in
+    thy
+    |> Typedef.add_typedef_global false (SOME raw_tyco) (raw_tyco, vs, NoSyn)
+        (HOLogic.mk_UNIV repT) NONE tac
+    |-> (fn (tyco, info) => get_typedef_info tyco vs info)
   end;
 
 fun mk_cons_tuple (left, right) =
@@ -127,7 +135,7 @@
   let
     val repT = HOLogic.mk_prodT (leftT, rightT);
 
-    val (((rep_inject, abs_inverse), absC, absT), typ_thy) =
+    val ((_, ((rep_inject, abs_inverse), absC, absT)), typ_thy) =
       thy
       |> do_typedef b repT alphas
       ||> Sign.add_path (Binding.name_of b); (*FIXME proper prefixing instead*)
@@ -155,7 +163,6 @@
       cdef_thy
       |> Iso_Tuple_Thms.map (Symtab.insert Thm.eq_thm_prop (isom_name, iso_tuple))
       |> Sign.restore_naming thy
-      |> Code.add_default_eqn isom_def;
   in
     ((isom, cons $ isom), thm_thy)
   end;
@@ -230,8 +237,8 @@
 val wN = "w";
 val moreN = "more";
 val schemeN = "_scheme";
-val ext_typeN = "_ext_type";
-val inner_typeN = "_inner_type";
+val ext_typeN = "_ext";
+val inner_typeN = "_inner";
 val extN ="_ext";
 val updateN = "_update";
 val makeN = "make";
@@ -1614,7 +1621,8 @@
 
     val ext_binding = Binding.name (suffix extN base_name);
     val ext_name = suffix extN name;
-    val extT = Type (suffix ext_typeN name, map TFree alphas_zeta);
+    val ext_tyco = suffix ext_typeN name
+    val extT = Type (ext_tyco, map TFree alphas_zeta);
     val ext_type = fields_moreTs ---> extT;
 
 
@@ -1768,10 +1776,9 @@
            [("ext_induct", induct),
             ("ext_inject", inject),
             ("ext_surjective", surject),
-            ("ext_split", split_meta)])
-      ||> Code.add_default_eqn ext_def;
-
-  in ((extT, induct', inject', surjective', split_meta', ext_def), thm_thy) end;
+            ("ext_split", split_meta)]);
+
+  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]
@@ -1815,6 +1822,73 @@
   in Thm.implies_elim thm' thm end;
 
 
+(* code generation *)
+
+fun instantiate_random_record tyco vs extN Ts thy =
+  let
+    val size = @{term "i::code_numeral"};
+    fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit => term"});
+    val T = Type (tyco, map TFree vs);
+    val Tm = termifyT T;
+    val params = Name.names Name.context "x" Ts;
+    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)
+        tc @{typ Random.seed} (SOME Tm, @{typ Random.seed});
+    val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
+  in 
+    thy
+    |> Class.instantiation ([tyco], vs, @{sort random})
+    |> `(fn lthy => Syntax.check_term lthy eq)
+    |-> (fn eq => Specification.definition (NONE, (apfst Binding.conceal Attrib.empty_binding, eq)))
+    |> snd
+    |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
+  end;
+
+fun ensure_random_record ext_tyco vs extN Ts thy =
+  let
+    val algebra = Sign.classes_of thy;
+    val has_inst = can (Sorts.mg_domain algebra ext_tyco) @{sort random};
+  in if has_inst then thy
+    else case Quickcheck_Generators.perhaps_constrain thy (map (rpair @{sort random}) Ts) vs
+     of SOME constrain => instantiate_random_record ext_tyco (map constrain vs) extN
+          ((map o map_atyps) (fn TFree v => TFree (constrain v)) Ts) thy
+      | NONE => thy
+  end;
+
+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 []
+      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}
+      |> Thm.instantiate
+         ([pairself (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort eq}), 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])
+    |> `(fn lthy => Syntax.check_term lthy eq)
+    |-> (fn eq => Specification.definition
+         (NONE, (Attrib.empty_binding, eq)))
+    |-> (fn (_, (_, eq_def)) =>
+       Class.prove_instantiation_exit_result Morphism.thm
+         (fn _ => fn eq_def => tac eq_def) eq_def)
+    |-> (fn eq_def => fn thy => thy |> Code.del_eqn eq_def |> Code.add_default_eqn (mk_eq thy eq_def))
+    |> (fn thy => Code.add_nbe_default_eqn (mk_eq_refl thy) thy)
+    |> ensure_random_record ext_tyco vs (fst ext) ((fst o strip_type o snd) ext)
+  end;
+
+
 (* definition *)
 
 fun definition (alphas, binding) parent (parents: parent_info list) raw_fields thy =
@@ -1870,7 +1944,7 @@
 
     val extension_name = full binding;
 
-    val ((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;
@@ -2042,12 +2116,6 @@
         upd_specs
       ||>> (PureThy.add_defs false o map (Thm.no_attributes o apfst (Binding.conceal o Binding.name)))
         [make_spec, fields_spec, extend_spec, truncate_spec]
-      |->
-        (fn defs as ((sel_defs, upd_defs), derived_defs) =>
-          fold Code.add_default_eqn sel_defs
-          #> fold Code.add_default_eqn upd_defs
-          #> fold Code.add_default_eqn derived_defs
-          #> pair defs)
       ||> Theory.checkpoint
     val (((sel_defs, upd_defs), derived_defs), defs_thy) =
       timeit_msg "record trfuns/tyabbrs/selectors/updates/make/fields/extend/truncate defs:"
@@ -2330,6 +2398,7 @@
       |> add_splits extension_id (split_meta', split_object', split_ex', induct_scheme')
       |> add_extfields extension_name (fields @ [(full_moreN, moreT)])
       |> add_fieldext (extension_name, snd extension) (names @ [full_moreN])
+      |> add_code ext_tyco vs extT ext simps' ext_inject
       |> Sign.restore_naming thy;
 
   in final_thy end;
--- a/src/HOL/Tools/refute.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/refute.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -658,7 +658,7 @@
       | Const (@{const_name Finite_Set.card}, _) => t
       | Const (@{const_name Finite_Set.finite}, _) => t
       | Const (@{const_name Orderings.less}, Type ("fun", [@{typ nat},
-        Type ("fun", [@{typ nat}, Type ("bool", [])])])) => t
+        Type ("fun", [@{typ nat}, @{typ bool}])])) => t
       | Const (@{const_name Groups.plus}, Type ("fun", [@{typ nat},
         Type ("fun", [@{typ nat}, @{typ nat}])])) => t
       | Const (@{const_name Groups.minus}, Type ("fun", [@{typ nat},
@@ -835,7 +835,7 @@
       | Const (@{const_name Finite_Set.finite}, T) =>
         collect_type_axioms T axs
       | Const (@{const_name Orderings.less}, T as Type ("fun", [@{typ nat},
-        Type ("fun", [@{typ nat}, Type ("bool", [])])])) =>
+        Type ("fun", [@{typ nat}, @{typ bool}])])) =>
           collect_type_axioms T axs
       | Const (@{const_name Groups.plus}, T as Type ("fun", [@{typ nat},
         Type ("fun", [@{typ nat}, @{typ nat}])])) =>
@@ -2653,7 +2653,7 @@
   fun Finite_Set_card_interpreter thy model args t =
     case t of
       Const (@{const_name Finite_Set.card},
-        Type ("fun", [Type ("fun", [T, Type ("bool", [])]),
+        Type ("fun", [Type ("fun", [T, @{typ bool}]),
                       @{typ nat}])) =>
       let
         (* interpretation -> int *)
@@ -2685,7 +2685,7 @@
               Leaf (replicate size_of_nat False)
           end
         val set_constants =
-          make_constants thy model (Type ("fun", [T, Type ("bool", [])]))
+          make_constants thy model (Type ("fun", [T, HOLogic.boolT]))
       in
         SOME (Node (map card set_constants), model, args)
       end
@@ -2702,17 +2702,17 @@
   fun Finite_Set_finite_interpreter thy model args t =
     case t of
       Const (@{const_name Finite_Set.finite},
-        Type ("fun", [Type ("fun", [T, Type ("bool", [])]),
-                      Type ("bool", [])])) $ _ =>
+        Type ("fun", [Type ("fun", [T, @{typ bool}]),
+                      @{typ bool}])) $ _ =>
         (* we only consider finite models anyway, hence EVERY set is *)
         (* "finite"                                                  *)
         SOME (TT, model, args)
     | Const (@{const_name Finite_Set.finite},
-        Type ("fun", [Type ("fun", [T, Type ("bool", [])]),
-                      Type ("bool", [])])) =>
+        Type ("fun", [Type ("fun", [T, @{typ bool}]),
+                      @{typ bool}])) =>
       let
         val size_of_set =
-          size_of_type thy model (Type ("fun", [T, Type ("bool", [])]))
+          size_of_type thy model (Type ("fun", [T, HOLogic.boolT]))
       in
         (* we only consider finite models anyway, hence EVERY set is *)
         (* "finite"                                                  *)
@@ -2731,7 +2731,7 @@
   fun Nat_less_interpreter thy model args t =
     case t of
       Const (@{const_name Orderings.less}, Type ("fun", [@{typ nat},
-        Type ("fun", [@{typ nat}, Type ("bool", [])])])) =>
+        Type ("fun", [@{typ nat}, @{typ bool}])])) =>
       let
         val size_of_nat = size_of_type thy model (@{typ nat})
         (* the 'n'-th nat is not less than the first 'n' nats, while it *)
@@ -2940,20 +2940,20 @@
   fun lfp_interpreter thy model args t =
     case t of
       Const (@{const_name lfp}, Type ("fun", [Type ("fun",
-        [Type ("fun", [T, Type ("bool", [])]),
-         Type ("fun", [_, Type ("bool", [])])]),
-         Type ("fun", [_, Type ("bool", [])])])) =>
+        [Type ("fun", [T, @{typ bool}]),
+         Type ("fun", [_, @{typ bool}])]),
+         Type ("fun", [_, @{typ bool}])])) =>
       let
         val size_elem = size_of_type thy model T
         (* the universe (i.e. the set that contains every element) *)
         val i_univ = Node (replicate size_elem TT)
         (* all sets with elements from type 'T' *)
         val i_sets =
-          make_constants thy model (Type ("fun", [T, Type ("bool", [])]))
+          make_constants thy model (Type ("fun", [T, HOLogic.boolT]))
         (* all functions that map sets to sets *)
         val i_funs = make_constants thy model (Type ("fun",
-          [Type ("fun", [T, Type ("bool", [])]),
-           Type ("fun", [T, Type ("bool", [])])]))
+          [Type ("fun", [T, @{typ bool}]),
+           Type ("fun", [T, @{typ bool}])]))
         (* "lfp(f) == Inter({u. f(u) <= u})" *)
         (* interpretation * interpretation -> bool *)
         fun is_subset (Node subs, Node sups) =
@@ -2995,20 +2995,20 @@
   fun gfp_interpreter thy model args t =
     case t of
       Const (@{const_name gfp}, Type ("fun", [Type ("fun",
-        [Type ("fun", [T, Type ("bool", [])]),
-         Type ("fun", [_, Type ("bool", [])])]),
-         Type ("fun", [_, Type ("bool", [])])])) =>
+        [Type ("fun", [T, @{typ bool}]),
+         Type ("fun", [_, @{typ bool}])]),
+         Type ("fun", [_, @{typ bool}])])) =>
       let
         val size_elem = size_of_type thy model T
         (* the universe (i.e. the set that contains every element) *)
         val i_univ = Node (replicate size_elem TT)
         (* all sets with elements from type 'T' *)
         val i_sets =
-          make_constants thy model (Type ("fun", [T, Type ("bool", [])]))
+          make_constants thy model (Type ("fun", [T, HOLogic.boolT]))
         (* all functions that map sets to sets *)
         val i_funs = make_constants thy model (Type ("fun",
-          [Type ("fun", [T, Type ("bool", [])]),
-           Type ("fun", [T, Type ("bool", [])])]))
+          [Type ("fun", [T, HOLogic.boolT]),
+           Type ("fun", [T, HOLogic.boolT])]))
         (* "gfp(f) == Union({u. u <= f(u)})" *)
         (* interpretation * interpretation -> bool *)
         fun is_subset (Node subs, Node sups) =
--- a/src/HOL/Tools/sat_funcs.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/sat_funcs.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -217,7 +217,7 @@
 	(* Thm.cterm -> int option *)
 	fun index_of_literal chyp = (
 		case (HOLogic.dest_Trueprop o Thm.term_of) chyp of
-		  (Const ("Not", _) $ atom) =>
+		  (Const (@{const_name Not}, _) $ atom) =>
 			SOME (~ (the (Termtab.lookup atom_table atom)))
 		| atom =>
 			SOME (the (Termtab.lookup atom_table atom))
--- a/src/HOL/Tools/simpdata.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Tools/simpdata.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -45,7 +45,7 @@
   (*expects Trueprop if not == *)
   of Const (@{const_name "=="},_) $ _ $ _ => th
    | _ $ (Const (@{const_name "op ="}, _) $ _ $ _) => mk_meta_eq th
-   | _ $ (Const (@{const_name "Not"}, _) $ _) => th RS @{thm Eq_FalseI}
+   | _ $ (Const (@{const_name Not}, _) $ _) => th RS @{thm Eq_FalseI}
    | _ => th RS @{thm Eq_TrueI}
 
 fun mk_eq_True (_: simpset) r =
--- a/src/HOL/Tools/typecopy.ML	Thu Aug 19 17:40:44 2010 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,134 +0,0 @@
-(*  Title:      HOL/Tools/typecopy.ML
-    Author:     Florian Haftmann, TU Muenchen
-
-Introducing copies of types using trivial typedefs; datatype-like abstraction.
-*)
-
-signature TYPECOPY =
-sig
-  type info = { vs: (string * sort) list, typ: typ, constr: string, proj: string,
-    constr_inject: thm, proj_inject: thm, constr_proj: thm, proj_constr: thm }
-  val typecopy: binding * (string * sort) list -> typ -> binding -> binding
-    -> theory -> (string * info) * theory
-  val get_info: theory -> string -> info option
-  val interpretation: (string -> theory -> theory) -> theory -> theory
-  val add_default_code: string -> theory -> theory
-  val setup: theory -> theory
-end;
-
-structure Typecopy: TYPECOPY =
-struct
-
-(* theory data *)
-
-type info = {
-  vs: (string * sort) list,
-  typ: typ,
-  constr: string,
-  proj: string,
-  constr_inject: thm,
-  proj_inject: thm,
-  constr_proj: thm,
-  proj_constr: thm
-};
-
-structure Typecopy_Data = Theory_Data
-(
-  type T = info Symtab.table;
-  val empty = Symtab.empty;
-  val extend = I;
-  fun merge data = Symtab.merge (K true) data;
-);
-
-val get_info = Symtab.lookup o Typecopy_Data.get;
-
-
-(* interpretation of type copies *)
-
-structure Typecopy_Interpretation = Interpretation(type T = string val eq = op =);
-val interpretation = Typecopy_Interpretation.interpretation;
-
-
-(* introducing typecopies *)
-
-fun add_info tyco vs (({ rep_type, Abs_name, Rep_name, ...},
-    { Abs_inject, Rep_inject, Abs_inverse, Rep_inverse, ... }) : Typedef.info) thy =
-  let
-    val exists_thm =
-      UNIV_I
-      |> Drule.instantiate' [SOME (ctyp_of thy (Logic.varifyT_global rep_type))] [];
-    val constr_inject = Abs_inject OF [exists_thm, exists_thm];
-    val proj_constr = Abs_inverse OF [exists_thm];
-    val info = {
-      vs = vs,
-      typ = rep_type,
-      constr = Abs_name,
-      proj = Rep_name,
-      constr_inject = constr_inject,
-      proj_inject = Rep_inject,
-      constr_proj = Rep_inverse,
-      proj_constr = proj_constr
-    };
-  in
-    thy
-    |> (Typecopy_Data.map o Symtab.update_new) (tyco, info)
-    |> Typecopy_Interpretation.data tyco
-    |> pair (tyco, info)
-  end
-
-fun typecopy (raw_tyco, raw_vs) raw_ty constr proj thy =
-  let
-    val ty = Sign.certify_typ thy raw_ty;
-    val ctxt = ProofContext.init_global thy |> Variable.declare_typ ty;
-    val vs = map (ProofContext.check_tfree ctxt) raw_vs;
-    val tac = Tactic.rtac UNIV_witness 1;
-  in
-    thy
-    |> Typedef.add_typedef_global false (SOME raw_tyco) (raw_tyco, vs, NoSyn)
-      (HOLogic.mk_UNIV ty) (SOME (proj, constr)) tac
-    |-> (fn (tyco, info) => add_info tyco vs info)
-  end;
-
-
-(* default code setup *)
-
-fun add_default_code tyco thy =
-  let
-    val SOME { vs, typ = ty_rep, constr = c, proj, proj_constr, proj_inject, ... } =
-      get_info thy tyco;
-    val constr = (c, Logic.unvarifyT_global (Sign.the_const_type thy c));
-    val ty = Type (tyco, map TFree vs);
-    val proj = Const (proj, ty --> ty_rep);
-    val (t_x, t_y) = (Free ("x", ty), Free ("y", ty));
-    val eq_lhs = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
-      $ t_x $ t_y;
-    val eq_rhs = HOLogic.mk_eq (proj $ t_x, proj $ t_y);
-    val eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq) (eq_lhs, eq_rhs);
-    fun tac eq_thm = Class.intro_classes_tac []
-      THEN (Simplifier.rewrite_goals_tac
-        (map Simpdata.mk_eq [eq_thm, @{thm eq}, proj_inject]))
-          THEN ALLGOALS (rtac @{thm refl});
-    fun mk_eq_refl thy = @{thm HOL.eq_refl}
-      |> Thm.instantiate
-         ([pairself (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort eq}), Logic.varifyT_global ty)], [])
-      |> AxClass.unoverload thy;
-  in
-    thy
-    |> Code.add_datatype [constr]
-    |> Code.add_eqn proj_constr
-    |> Class.instantiation ([tyco], vs, [HOLogic.class_eq])
-    |> `(fn lthy => Syntax.check_term lthy eq)
-    |-> (fn eq => Specification.definition
-         (NONE, (Attrib.empty_binding, eq)))
-    |-> (fn (_, (_, eq_thm)) =>
-       Class.prove_instantiation_exit_result Morphism.thm
-         (fn _ => fn eq_thm => tac eq_thm) eq_thm)
-    |-> (fn eq_thm => Code.add_eqn eq_thm)
-    |> (fn thy => Code.add_nbe_eqn (mk_eq_refl thy) thy)
-  end;
-
-val setup =
-  Typecopy_Interpretation.init
-  #> interpretation add_default_code
-
-end;
--- a/src/HOL/Tools/typedef_codegen.ML	Thu Aug 19 17:40:44 2010 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,110 +0,0 @@
-(*  Title:      HOL/Tools/typedef_codegen.ML
-    Author:     Stefan Berghofer, TU Muenchen
-
-Code generators for trivial typedefs.
-*)
-
-signature TYPEDEF_CODEGEN =
-sig
-  val setup: theory -> theory
-end;
-
-structure TypedefCodegen: TYPEDEF_CODEGEN =
-struct
-
-fun typedef_codegen thy defs dep module brack t gr =
-  let
-    fun get_name (Type (tname, _)) = tname
-      | get_name _ = "";
-    fun mk_fun s T ts =
-      let
-        val (_, gr') = Codegen.invoke_tycodegen thy defs dep module false T gr;
-        val (ps, gr'') =
-          fold_map (Codegen.invoke_codegen thy defs dep module true) ts gr';
-        val id = Codegen.mk_qual_id module (Codegen.get_const_id gr'' s)
-      in SOME (Codegen.mk_app brack (Codegen.str id) ps, gr'') end;
-    fun lookup f T =
-      (* FIXME handle multiple typedef interpretations (!??) *)
-      (case Typedef.get_info_global thy (get_name T) of
-        [info] => f info
-      | _ => "");
-  in
-    (case strip_comb t of
-       (Const (s, Type ("fun", [T, U])), ts) =>
-         if lookup (#Rep_name o #1) T = s andalso
-           is_none (Codegen.get_assoc_type thy (get_name T))
-         then mk_fun s T ts
-         else if lookup (#Abs_name o #1) U = s andalso
-           is_none (Codegen.get_assoc_type thy (get_name U))
-         then mk_fun s U ts
-         else NONE
-     | _ => NONE)
-  end;
-
-fun mk_tyexpr [] s = Codegen.str s
-  | mk_tyexpr [p] s = Pretty.block [p, Codegen.str (" " ^ s)]
-  | mk_tyexpr ps s = Pretty.list "(" (") " ^ s) ps;
-
-fun typedef_tycodegen thy defs dep module brack (Type (s, Ts)) gr =
-      (case Typedef.get_info_global thy s of
-        (* FIXME handle multiple typedef interpretations (!??) *)
-        [({abs_type as newT as Type (tname, Us), rep_type = oldT, Abs_name, Rep_name, ...}, _)] =>
-          if is_some (Codegen.get_assoc_type thy tname) then NONE
-          else
-            let
-              val module' = Codegen.if_library
-                (Codegen.thyname_of_type thy tname) module;
-              val node_id = tname ^ " (type)";
-              val ((((qs, (_, Abs_id)), (_, Rep_id)), ty_id), gr') = gr |> fold_map
-                  (Codegen.invoke_tycodegen thy defs dep module (length Ts = 1))
-                    Ts ||>>
-                Codegen.mk_const_id module' Abs_name ||>>
-                Codegen.mk_const_id module' Rep_name ||>>
-                Codegen.mk_type_id module' s;
-              val tyexpr = mk_tyexpr qs (Codegen.mk_qual_id module ty_id)
-            in
-              SOME (tyexpr, case try (Codegen.get_node gr') node_id of
-                NONE =>
-                  let
-                    val (p :: ps, gr'') = fold_map
-                      (Codegen.invoke_tycodegen thy defs node_id module' false)
-                      (oldT :: Us) (Codegen.add_edge (node_id, dep)
-                        (Codegen.new_node (node_id, (NONE, "", "")) gr'));
-                    val s =
-                      Codegen.string_of (Pretty.block [Codegen.str "datatype ",
-                        mk_tyexpr ps (snd ty_id),
-                        Codegen.str " =", Pretty.brk 1, Codegen.str (Abs_id ^ " of"),
-                        Pretty.brk 1, p, Codegen.str ";"]) ^ "\n\n" ^
-                      Codegen.string_of (Pretty.block [Codegen.str ("fun " ^ Rep_id),
-                        Pretty.brk 1, Codegen.str ("(" ^ Abs_id), Pretty.brk 1,
-                        Codegen.str "x) = x;"]) ^ "\n\n" ^
-                      (if member (op =) (!Codegen.mode) "term_of" then
-                        Codegen.string_of (Pretty.block [Codegen.str "fun ",
-                          Codegen.mk_term_of gr'' module' false newT, Pretty.brk 1,
-                          Codegen.str ("(" ^ Abs_id), Pretty.brk 1,
-                          Codegen.str "x) =", Pretty.brk 1,
-                          Pretty.block [Codegen.str ("Const (\"" ^ Abs_name ^ "\","),
-                            Pretty.brk 1, Codegen.mk_type false (oldT --> newT),
-                            Codegen.str ")"], Codegen.str " $", Pretty.brk 1,
-                          Codegen.mk_term_of gr'' module' false oldT, Pretty.brk 1,
-                          Codegen.str "x;"]) ^ "\n\n"
-                       else "") ^
-                      (if member (op =) (!Codegen.mode) "test" then
-                        Codegen.string_of (Pretty.block [Codegen.str "fun ",
-                          Codegen.mk_gen gr'' module' false [] "" newT, Pretty.brk 1,
-                          Codegen.str "i =", Pretty.brk 1,
-                          Pretty.block [Codegen.str (Abs_id ^ " ("),
-                            Codegen.mk_gen gr'' module' false [] "" oldT, Pretty.brk 1,
-                            Codegen.str "i);"]]) ^ "\n\n"
-                       else "")
-                  in Codegen.map_node node_id (K (NONE, module', s)) gr'' end
-              | SOME _ => Codegen.add_edge (node_id, dep) gr')
-            end
-      | _ => NONE)
-  | typedef_tycodegen thy defs dep module brack _ gr = NONE;
-
-val setup =
-  Codegen.add_codegen "typedef" typedef_codegen
-  #> Codegen.add_tycodegen "typedef" typedef_tycodegen
-
-end;
--- a/src/HOL/Typedef.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Typedef.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -6,9 +6,7 @@
 
 theory Typedef
 imports Set
-uses
-  ("Tools/typedef.ML")
-  ("Tools/typedef_codegen.ML")
+uses ("Tools/typedef.ML")
 begin
 
 ML {*
@@ -115,6 +113,5 @@
 end
 
 use "Tools/typedef.ML" setup Typedef.setup
-use "Tools/typedef_codegen.ML" setup TypedefCodegen.setup
 
 end
--- a/src/HOL/Word/Word.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/Word/Word.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -4034,6 +4034,11 @@
   "length (rotater n xs) = length xs"
   by (simp add : rotater_rev)
 
+lemma restrict_to_left:
+  assumes "x = y"
+  shows "(x = z) = (y = z)"
+  using assms by simp
+
 lemmas rrs0 = rotate_eqs [THEN restrict_to_left, 
   simplified rotate_gal [symmetric] rotate_gal' [symmetric], standard]
 lemmas rrs1 = rrs0 [THEN refl [THEN rev_iffD1]]
--- a/src/HOL/ex/SAT_Examples.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/ex/SAT_Examples.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -273,16 +273,7 @@
 and 183: "~x29 | ~x58"
 and 184: "~x28 | ~x58"
 shows "False"
-using 1 2 3 4 5 6 7 8 9 10 11 12 13 14 15 16 17 18 19
-20 21 22 23 24 25 26 27 28 29 30 31 32 33 34 35 36 37 38 39
-40 41 42 43 44 45 46 47 48 49 50 51 52 53 54 55 56 57 58 59
-60 61 62 63 64 65 66 67 68 69 70 71 72 73 74 75 76 77 78 79
-80 81 82 83 84 85 86 87 88 89 90 91 92 93 94 95 96 97 98 99
-100 101 102 103 104 105 106 107 108 109 110 111 112 113 114 115 116 117 118 119
-120 121 122 123 124 125 126 127 128 129 130 131 132 133 134 135 136 137 138 139
-140 141 142 143 144 145 146 147 148 149 150 151 152 153 154 155 156 157 158 159
-160 161 162 163 164 165 166 167 168 169 170 171 172 173 174 175 176 177 178 179
-180 181 182 183 184
+using assms
 (*
 by sat
 *)
@@ -495,17 +486,7 @@
 and 203: "~x41 | ~x55"
 and 204: "~x48 | ~x55"
 shows "False"
-using 1 2 3 4 5 6 7 8 9 10 11 12 13 14 15 16 17 18 19
-20 21 22 23 24 25 26 27 28 29 30 31 32 33 34 35 36 37 38 39
-40 41 42 43 44 45 46 47 48 49 50 51 52 53 54 55 56 57 58 59
-60 61 62 63 64 65 66 67 68 69 70 71 72 73 74 75 76 77 78 79
-80 81 82 83 84 85 86 87 88 89 90 91 92 93 94 95 96 97 98 99
-100 101 102 103 104 105 106 107 108 109 110 111 112 113 114 115 116 117 118 119
-120 121 122 123 124 125 126 127 128 129 130 131 132 133 134 135 136 137 138 139
-140 141 142 143 144 145 146 147 148 149 150 151 152 153 154 155 156 157 158 159
-160 161 162 163 164 165 166 167 168 169 170 171 172 173 174 175 176 177 178 179
-180 181 182 183 184 185 186 187 188 189 190 191 192 193 194 195 196 197 198 199
-200 201 202 203 204
+using assms
 (*
 by sat
 *)
--- a/src/HOL/ex/SVC_Oracle.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/ex/SVC_Oracle.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -88,18 +88,18 @@
             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("op &", _)) $ p $ q) = c $ (fm p) $ (fm q)
-      | fm ((c as Const("op |", _)) $ p $ q) = c $ (fm p) $ (fm q)
-      | fm ((c as Const("op -->", _)) $ p $ q) = c $ (fm p) $ (fm q)
-      | fm ((c as Const("Not", _)) $ p) = c $ (fm p)
-      | fm ((c as Const("True", _))) = c
-      | fm ((c as Const("False", _))) = c
-      | fm (t as Const("op =",  Type ("fun", [T,_])) $ _ $ _) = rel (T, t)
+    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)
+      | 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 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
     (*entry point, and abstraction of a meta-formula*)
-    fun mt ((c as Const("Trueprop", _)) $ p) = c $ (fm p)
+    fun mt ((c as Const(@{const_name Trueprop}, _)) $ p) = c $ (fm p)
       | mt ((c as Const("==>", _)) $ p $ q)  = c $ (mt p) $ (mt q)
       | mt t = fm t  (*it might be a formula*)
   in (list_all (params, mt body), !pairs) end;
--- a/src/HOL/ex/svc_funcs.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOL/ex/svc_funcs.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -89,13 +89,13 @@
    let fun tag t =
          let val (c,ts) = strip_comb t
          in  case c of
-             Const("op &", _)   => apply c (map tag ts)
-           | Const("op |", _)   => apply c (map tag ts)
-           | Const("op -->", _) => apply c (map tag ts)
-           | Const("Not", _)    => apply c (map tag ts)
-           | Const("True", _)   => (c, false)
-           | Const("False", _)  => (c, false)
-           | Const("op =", Type ("fun", [T,_])) =>
+             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 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,_])) =>
                  if T = HOLogic.boolT then
                      (*biconditional: with int/nat comparisons below?*)
                      let val [t1,t2] = ts
@@ -183,16 +183,16 @@
       | tm t = Int (lit t)
                handle Match => var (t,[])
     (*translation of a formula*)
-    and fm pos (Const("op &", _) $ p $ q) =
+    and fm pos (Const(@{const_name "op &"}, _) $ p $ q) =
             Buildin("AND", [fm pos p, fm pos q])
-      | fm pos (Const("op |", _) $ p $ q) =
+      | fm pos (Const(@{const_name "op |"}, _) $ p $ q) =
             Buildin("OR", [fm pos p, fm pos q])
-      | fm pos (Const("op -->", _) $ p $ q) =
+      | fm pos (Const(@{const_name "op -->"}, _) $ p $ q) =
             Buildin("=>", [fm (not pos) p, fm pos q])
-      | fm pos (Const("Not", _) $ p) =
+      | fm pos (Const(@{const_name Not}, _) $ p) =
             Buildin("NOT", [fm (not pos) p])
-      | fm pos (Const("True", _)) = TrueExpr
-      | fm pos (Const("False", _)) = FalseExpr
+      | fm pos (Const(@{const_name True}, _)) = TrueExpr
+      | fm pos (Const(@{const_name False}, _)) = FalseExpr
       | fm pos (Const("SVC_Oracle.iff_keep", _) $ p $ q) =
              (*polarity doesn't matter*)
             Buildin("=", [fm pos p, fm pos q])
@@ -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("op =", Type ("fun", [T,_])) $ x $ y) =
+      | fm pos (t as Const(@{const_name "op ="}, 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])
@@ -225,7 +225,7 @@
             else fail t
       | fm pos t = var(t,[]);
       (*entry point, and translation of a meta-formula*)
-      fun mt pos ((c as Const("Trueprop", _)) $ p) = fm pos (iff_tag p)
+      fun mt pos ((c as Const(@{const_name Trueprop}, _)) $ p) = fm pos (iff_tag p)
         | mt pos ((c as Const("==>", _)) $ p $ q) =
             Buildin("=>", [mt (not pos) p, mt pos q])
         | mt pos t = fm pos (iff_tag t)  (*it might be a formula*)
--- a/src/HOLCF/Tools/Domain/domain_library.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/HOLCF/Tools/Domain/domain_library.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -121,8 +121,8 @@
       val r_inst = read_instantiate ctxt;
       fun at thm =
           case concl_of thm of
-            _$(Const("op &",_)$_$_)       => at(thm RS conjunct1)@at(thm RS conjunct2)
-          | _$(Const("All" ,_)$Abs(s,_,_))=> at(thm RS (r_inst [(("x", 0), "?" ^ s)] spec))
+            _$(Const(@{const_name "op &"},_)$_$_)       => 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;
 
@@ -185,7 +185,7 @@
 fun mk_lam  (x,T) = Abs(x,dummyT,T);
 fun mk_all  (x,P) = HOLogic.mk_all (x,dummyT,P);
 fun mk_ex   (x,P) = mk_exists (x,dummyT,P);
-fun mk_constrainall (x,typ,P) = %%:"All" $ (Type_Infer.constrain (typ --> boolT) (mk_lam(x,P)));
+fun mk_constrainall (x,typ,P) = %%: @{const_name All} $ (Type_Infer.constrain (typ --> boolT) (mk_lam(x,P)));
 end
 
 fun mk_All  (x,P) = %%:"all" $ mk_lam(x,P); (* meta universal quantification *)
--- a/src/Pure/Isar/class_declaration.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/Pure/Isar/class_declaration.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -129,18 +129,22 @@
                   ^ 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) =>
+        if Type_Infer.is_param vi then Type_Infer.param 0 (Name.aT, sort) else T);
     fun add_typ_check level name f = Context.proof_map
       (Syntax.add_typ_check level name (fn Ts => fn ctxt =>
         let val Ts' = f Ts in if eq_list (op =) (Ts, Ts') then NONE else SOME (Ts', ctxt) end));
 
     (* preprocessing elements, retrieving base sort from type-checked elements *)
+    val raw_supexpr = (map (fn sup => (sup, (("", false),
+      Expression.Positional []))) sups, []);
     val init_class_body = fold (ProofContext.add_const_constraint o apsnd SOME) base_constraints
       #> Class.redeclare_operations thy sups
       #> add_typ_check 10 "reject_bcd_etc" reject_bcd_etc
       #> add_typ_check ~10 "singleton_fixate" singleton_fixate;
-    val raw_supexpr = (map (fn sup => (sup, (("", false),
-      Expression.Positional []))) sups, []);
     val ((raw_supparams, _, raw_inferred_elems), _) = ProofContext.init_global thy
+      |> add_typ_check 5 "after_infer_fixate" after_infer_fixate
       |> prep_decl raw_supexpr init_class_body raw_elems;
     fun filter_element (Element.Fixes []) = NONE
       | filter_element (e as Element.Fixes _) = SOME e
--- a/src/Sequents/LK0.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/Sequents/LK0.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -12,8 +12,6 @@
 imports Sequents
 begin
 
-global
-
 classes "term"
 default_sort "term"
 
@@ -61,8 +59,6 @@
   Ex  (binder "\<exists>" 10) and
   not_equal  (infixl "\<noteq>" 50)
 
-local
-
 axioms
 
   (*Structural rules: contraction, thinning, exchange [Soren Heilmann] *)
--- a/src/Sequents/Sequents.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/Sequents/Sequents.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -14,8 +14,6 @@
 
 declare [[unify_trace_bound = 20, unify_search_bound = 40]]
 
-global
-
 typedecl o
 
 (* Sequences *)
--- a/src/Sequents/modal.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/Sequents/modal.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -29,7 +29,7 @@
 in 
 
 (*Returns the list of all formulas in the sequent*)
-fun forms_of_seq (Const("SeqO",_) $ P $ u) = P :: forms_of_seq u
+fun forms_of_seq (Const(@{const_name SeqO'},_) $ P $ u) = P :: forms_of_seq u
   | forms_of_seq (H $ u) = forms_of_seq u
   | forms_of_seq _ = [];
 
--- a/src/Sequents/prover.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/Sequents/prover.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -55,7 +55,7 @@
     ["Unsafe rules:"] @ map Display.string_of_thm_without_context unsafes));
 
 (*Returns the list of all formulas in the sequent*)
-fun forms_of_seq (Const("SeqO'",_) $ P $ u) = P :: forms_of_seq u
+fun forms_of_seq (Const(@{const_name "SeqO'"},_) $ P $ u) = P :: forms_of_seq u
   | forms_of_seq (H $ u) = forms_of_seq u
   | forms_of_seq _ = [];
 
--- a/src/Sequents/simpdata.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/Sequents/simpdata.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -13,16 +13,16 @@
 (*Make atomic rewrite rules*)
 fun atomize r =
  case concl_of r of
-   Const("Trueprop",_) $ Abs(_,_,a) $ Abs(_,_,c) =>
+   Const(@{const_name Trueprop},_) $ Abs(_,_,a) $ Abs(_,_,c) =>
      (case (forms_of_seq a, forms_of_seq c) of
         ([], [p]) =>
           (case p of
-               Const("imp",_)$_$_ => atomize(r RS @{thm mp_R})
-             | Const("conj",_)$_$_   => atomize(r RS @{thm conjunct1}) @
+               Const(@{const_name imp},_)$_$_ => atomize(r RS @{thm mp_R})
+             | Const(@{const_name conj},_)$_$_   => atomize(r RS @{thm conjunct1}) @
                    atomize(r RS @{thm conjunct2})
-             | Const("All",_)$_      => atomize(r RS @{thm spec})
-             | Const("True",_)       => []    (*True is DELETED*)
-             | Const("False",_)      => []    (*should False do something?*)
+             | Const(@{const_name All},_)$_      => atomize(r RS @{thm spec})
+             | Const(@{const_name True},_)       => []    (*True is DELETED*)
+             | Const(@{const_name False},_)      => []    (*should False do something?*)
              | _                     => [r])
       | _ => [])  (*ignore theorem unless it has precisely one conclusion*)
  | _ => [r];
@@ -30,13 +30,13 @@
 (*Make meta-equalities.*)
 fun mk_meta_eq th = case concl_of th of
     Const("==",_)$_$_           => th
-  | Const("Trueprop",_) $ Abs(_,_,a) $ Abs(_,_,c) =>
+  | Const(@{const_name Trueprop},_) $ Abs(_,_,a) $ Abs(_,_,c) =>
         (case (forms_of_seq a, forms_of_seq c) of
              ([], [p]) =>
                  (case p of
-                      (Const("equal",_)$_$_)   => th RS @{thm eq_reflection}
-                    | (Const("iff",_)$_$_) => th RS @{thm iff_reflection}
-                    | (Const("Not",_)$_)      => th RS @{thm iff_reflection_F}
+                      (Const(@{const_name equal},_)$_$_)   => th RS @{thm eq_reflection}
+                    | (Const(@{const_name iff},_)$_$_) => th RS @{thm iff_reflection}
+                    | (Const(@{const_name Not},_)$_)      => th RS @{thm iff_reflection_F}
                     | _                       => th RS @{thm iff_reflection_T})
            | _ => error ("addsimps: unable to use theorem\n" ^
                          Display.string_of_thm_without_context th));
--- a/src/ZF/Datatype_ZF.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/ZF/Datatype_ZF.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -61,7 +61,7 @@
 struct
   val trace = Unsynchronized.ref false;
 
-  fun mk_new ([],[]) = Const("True",FOLogic.oT)
+  fun mk_new ([],[]) = Const(@{const_name True},FOLogic.oT)
     | mk_new (largs,rargs) =
         Balanced_Tree.make FOLogic.mk_conj
                  (map FOLogic.mk_eq (ListPair.zip (largs,rargs)));
@@ -85,7 +85,7 @@
            if #big_rec_name lcon_info = #big_rec_name rcon_info
                andalso not (null (#free_iffs lcon_info)) then
                if lname = rname then mk_new (largs, rargs)
-               else Const("False",FOLogic.oT)
+               else Const(@{const_name False},FOLogic.oT)
            else raise Match
        val _ =
          if !trace then writeln ("NEW = " ^ Syntax.string_of_term_global sg new)
--- a/src/ZF/Inductive_ZF.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/ZF/Inductive_ZF.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -31,8 +31,8 @@
 lemma refl_thin: "!!P. a = a ==> P ==> P" .
 
 use "ind_syntax.ML"
+use "Tools/ind_cases.ML"
 use "Tools/cartprod.ML"
-use "Tools/ind_cases.ML"
 use "Tools/inductive_package.ML"
 use "Tools/induct_tacs.ML"
 use "Tools/primrec_package.ML"
--- a/src/ZF/Sum.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/ZF/Sum.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -9,8 +9,6 @@
 
 text{*And the "Part" primitive for simultaneous recursive type definitions*}
 
-global
-
 definition sum :: "[i,i]=>i" (infixr "+" 65) where
      "A+B == {0}*A Un {1}*B"
 
@@ -27,8 +25,6 @@
 definition Part :: "[i,i=>i] => i" where
      "Part(A,h) == {x: A. EX z. x = h(z)}"
 
-local
-
 subsection{*Rules for the @{term Part} Primitive*}
 
 lemma Part_iff: 
--- a/src/ZF/Tools/induct_tacs.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/ZF/Tools/induct_tacs.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -118,7 +118,7 @@
 fun rep_datatype_i elim induct case_eqns recursor_eqns thy =
   let
     (*analyze the LHS of a case equation to get a constructor*)
-    fun const_of (Const("op =", _) $ (_ $ c) $ _) = c
+    fun const_of (Const(@{const_name "op ="}, _) $ (_ $ c) $ _) = c
       | const_of eqn = error ("Ill-formed case equation: " ^
                               Syntax.string_of_term_global thy eqn);
 
--- a/src/ZF/Tools/inductive_package.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/ZF/Tools/inductive_package.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -102,7 +102,7 @@
 
   val z' = mk_variant"z" and X' = mk_variant"X" and w' = mk_variant"w";
 
-  fun dest_tprop (Const("Trueprop",_) $ P) = P
+  fun dest_tprop (Const(@{const_name Trueprop},_) $ P) = P
     | dest_tprop Q = error ("Ill-formed premise of introduction rule: " ^
                             Syntax.string_of_term ctxt Q);
 
--- a/src/ZF/Tools/primrec_package.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/ZF/Tools/primrec_package.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -115,8 +115,8 @@
             case AList.lookup (op =) eqns cname of
                 NONE => (warning ("no equation for constructor " ^ cname ^
                                   "\nin definition of function " ^ fname);
-                         (Const ("0", Ind_Syntax.iT),
-                          #2 recursor_pair, Const ("0", Ind_Syntax.iT)))
+                         (Const (@{const_name 0}, Ind_Syntax.iT),
+                          #2 recursor_pair, Const (@{const_name 0}, Ind_Syntax.iT)))
               | SOME (rhs, cargs', eq) =>
                     (rhs, inst_recursor (recursor_pair, cargs'), eq)
           val allowed_terms = map use_fabs (#2 (strip_comb recursor_rhs))
--- a/src/ZF/Tools/typechk.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/ZF/Tools/typechk.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -75,7 +75,7 @@
          if length rls <= maxr then resolve_tac rls i else no_tac
       end);
 
-fun is_rigid_elem (Const("Trueprop",_) $ (Const(@{const_name mem},_) $ a $ _)) =
+fun is_rigid_elem (Const(@{const_name Trueprop},_) $ (Const(@{const_name mem},_) $ a $ _)) =
       not (is_Var (head_of a))
   | is_rigid_elem _ = false;
 
--- a/src/ZF/ZF.thy	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/ZF/ZF.thy	Thu Aug 19 17:41:52 2010 +0200
@@ -12,8 +12,6 @@
 
 ML {* Unsynchronized.reset eta_contract *}
 
-global
-
 typedecl i
 arities  i :: "term"
 
@@ -209,8 +207,6 @@
   subset_def:    "A <= B == \<forall>x\<in>A. x\<in>B"
 
 
-local
-
 axioms
 
   (* ZF axioms -- see Suppes p.238
--- a/src/ZF/arith_data.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/ZF/arith_data.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -24,12 +24,12 @@
 
 val iT = Ind_Syntax.iT;
 
-val zero = Const("0", iT);
-val succ = Const("succ", iT --> iT);
+val zero = Const(@{const_name 0}, iT);
+val succ = Const(@{const_name succ}, iT --> iT);
 fun mk_succ t = succ $ t;
 val one = mk_succ zero;
 
-val mk_plus = FOLogic.mk_binop "Arith.add";
+val mk_plus = FOLogic.mk_binop @{const_name Arith.add};
 
 (*Thus mk_sum[t] yields t+#0; longer sums don't have a trailing zero*)
 fun mk_sum []        = zero
@@ -40,13 +40,13 @@
 fun long_mk_sum []        = zero
   | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
 
-val dest_plus = FOLogic.dest_bin "Arith.add" iT;
+val dest_plus = FOLogic.dest_bin @{const_name Arith.add} iT;
 
 (* dest_sum *)
 
-fun dest_sum (Const("0",_)) = []
-  | dest_sum (Const("succ",_) $ t) = one :: dest_sum t
-  | dest_sum (Const("Arith.add",_) $ t $ u) = dest_sum t @ dest_sum u
+fun dest_sum (Const(@{const_name 0},_)) = []
+  | dest_sum (Const(@{const_name succ},_) $ t) = one :: dest_sum t
+  | dest_sum (Const(@{const_name Arith.add},_) $ t $ u) = dest_sum t @ dest_sum u
   | dest_sum tm = [tm];
 
 (*Apply the given rewrite (if present) just once*)
@@ -83,14 +83,14 @@
 (*** Use CancelNumerals simproc without binary numerals,
      just for cancellation ***)
 
-val mk_times = FOLogic.mk_binop "Arith.mult";
+val mk_times = FOLogic.mk_binop @{const_name Arith.mult};
 
 fun mk_prod [] = one
   | mk_prod [t] = t
   | mk_prod (t :: ts) = if t = one then mk_prod ts
                         else mk_times (t, mk_prod ts);
 
-val dest_times = FOLogic.dest_bin "Arith.mult" iT;
+val dest_times = FOLogic.dest_bin @{const_name Arith.mult} iT;
 
 fun dest_prod t =
       let val (t,u) = dest_times t
@@ -174,8 +174,8 @@
   struct
   open CancelNumeralsCommon
   val prove_conv = prove_conv "natless_cancel_numerals"
-  val mk_bal   = FOLogic.mk_binrel "Ordinal.lt"
-  val dest_bal = FOLogic.dest_bin "Ordinal.lt" iT
+  val mk_bal   = FOLogic.mk_binrel @{const_name Ordinal.lt}
+  val dest_bal = FOLogic.dest_bin @{const_name Ordinal.lt} iT
   val bal_add1 = @{thm less_add_iff} RS @{thm iff_trans}
   val bal_add2 = @{thm less_add_iff} RS @{thm iff_trans}
   fun trans_tac _ = gen_trans_tac @{thm iff_trans}
@@ -187,8 +187,8 @@
   struct
   open CancelNumeralsCommon
   val prove_conv = prove_conv "natdiff_cancel_numerals"
-  val mk_bal   = FOLogic.mk_binop "Arith.diff"
-  val dest_bal = FOLogic.dest_bin "Arith.diff" iT
+  val mk_bal   = FOLogic.mk_binop @{const_name Arith.diff}
+  val dest_bal = FOLogic.dest_bin @{const_name Arith.diff} iT
   val bal_add1 = @{thm diff_add_eq} RS @{thm trans}
   val bal_add2 = @{thm diff_add_eq} RS @{thm trans}
   fun trans_tac _ = gen_trans_tac @{thm trans}
--- a/src/ZF/ind_syntax.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/ZF/ind_syntax.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -18,7 +18,7 @@
 
 (** Abstract syntax definitions for ZF **)
 
-val iT = Type("i",[]);
+val iT = Type(@{type_name i}, []);
 
 (*Creates All(%v.v:A --> P(v)) rather than Ball(A,P) *)
 fun mk_all_imp (A,P) =
--- a/src/ZF/simpdata.ML	Thu Aug 19 17:40:44 2010 +0200
+++ b/src/ZF/simpdata.ML	Thu Aug 19 17:41:52 2010 +0200
@@ -19,27 +19,27 @@
                    | NONE => [th])
             | _ => [th]
   in case concl_of th of
-         Const("Trueprop",_) $ P =>
+         Const(@{const_name Trueprop},_) $ P =>
             (case P of
                  Const(@{const_name mem},_) $ a $ b => tryrules mem_pairs b
-               | Const("True",_)         => []
-               | Const("False",_)        => []
+               | Const(@{const_name True},_)         => []
+               | Const(@{const_name False},_)        => []
                | A => tryrules conn_pairs A)
        | _                       => [th]
   end;
 
 (*Analyse a rigid formula*)
 val ZF_conn_pairs =
-  [("Ball",     [@{thm bspec}]),
-   ("All",      [@{thm spec}]),
-   ("op -->",   [@{thm mp}]),
-   ("op &",     [@{thm conjunct1}, @{thm conjunct2}])];
+  [(@{const_name Ball}, [@{thm bspec}]),
+   (@{const_name All}, [@{thm spec}]),
+   (@{const_name "op -->"}, [@{thm mp}]),
+   (@{const_name "op &"}, [@{thm conjunct1}, @{thm conjunct2}])];
 
 (*Analyse a:b, where b is rigid*)
 val ZF_mem_pairs =
-  [("Collect",  [@{thm CollectD1}, @{thm CollectD2}]),
-   (@{const_name Diff},     [@{thm DiffD1}, @{thm DiffD2}]),
-   (@{const_name Int},   [@{thm IntD1}, @{thm IntD2}])];
+  [(@{const_name Collect}, [@{thm CollectD1}, @{thm CollectD2}]),
+   (@{const_name Diff}, [@{thm DiffD1}, @{thm DiffD2}]),
+   (@{const_name Int}, [@{thm IntD1}, @{thm IntD2}])];
 
 val ZF_atomize = atomize (ZF_conn_pairs, ZF_mem_pairs);