merged
authornipkow
Wed, 10 Dec 2008 10:28:16 +0100
changeset 29027 501780b0bcae
parent 29025 8c8859c0d734 (diff)
parent 29026 5fbaa05f637f (current diff)
child 29036 df1113d916db
child 29106 25e28a4070f3
child 29677 6463bd4e5167
merged
src/HOL/ContNotDenum.thy
--- a/Admin/isatest/isatest-makeall	Wed Dec 10 10:23:47 2008 +0100
+++ b/Admin/isatest/isatest-makeall	Wed Dec 10 10:28:16 2008 +0100
@@ -133,6 +133,11 @@
 
     echo ------------------- starting test --- `date` --- $HOSTNAME > $TESTLOG 2>&1
 
+    if [ "${ISABELLE_HOME_USER:0:14}" == "/tmp/isabelle-" ]; then
+        echo "--- cleaning up old $ISABELLE_HOME_USER"
+        rm -rf $ISABELLE_HOME_USER
+    fi
+
     cat $SETTINGS >> $DISTPREFIX/Isabelle/etc/settings
     (ulimit -t $MAXTIME; cd $DIR; $NICE $TOOL >> $TESTLOG 2>&1)
 
--- a/Admin/isatest/settings/sun-poly	Wed Dec 10 10:23:47 2008 +0100
+++ b/Admin/isatest/settings/sun-poly	Wed Dec 10 10:28:16 2008 +0100
@@ -6,7 +6,7 @@
   ML_HOME="$POLYML_HOME/$ML_PLATFORM"
   ML_OPTIONS="-H 1500"
 
-ISABELLE_HOME_USER=~/isabelle-sun-poly
+ISABELLE_HOME_USER=/tmp/isabelle-sun-poly
 
 # Where to look for isabelle tools (multiple dirs separated by ':').
 ISABELLE_TOOLS="$ISABELLE_HOME/lib/Tools"
--- a/NEWS	Wed Dec 10 10:23:47 2008 +0100
+++ b/NEWS	Wed Dec 10 10:28:16 2008 +0100
@@ -61,6 +61,18 @@
 
 *** Pure ***
 
+* Type Binding.T gradually replaces formerly used type bstring for names
+to be bound.  Name space interface for declarations has been simplified:
+
+  NameSpace.declare: NameSpace.naming
+    -> Binding.T -> NameSpace.T -> string * NameSpace.T
+  NameSpace.bind: NameSpace.naming
+    -> Binding.T * 'a -> 'a NameSpace.table -> string * 'a NameSpace.table
+         (*exception Symtab.DUP*)
+
+See further modules src/Pure/General/binding.ML and
+src/Pure/General/name_space.ML
+
 * Module moves in repository:
     src/Pure/Tools/value.ML ~> src/Tools/
     src/Pure/Tools/quickcheck.ML ~> src/Tools/
--- a/doc-src/IsarAdvanced/Classes/classes.tex	Wed Dec 10 10:23:47 2008 +0100
+++ b/doc-src/IsarAdvanced/Classes/classes.tex	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,3 @@
-
-%% $Id$
 
 \documentclass[12pt,a4paper,fleqn]{report}
 \usepackage{latexsym,graphicx}
@@ -10,40 +8,10 @@
 \usepackage{../../pdfsetup}
 
 
-%% setup
-
-% hyphenation
 \hyphenation{Isabelle}
 \hyphenation{Isar}
-
-% logical markup
-\newcommand{\strong}[1]{{\bfseries {#1}}}
-\newcommand{\qn}[1]{\emph{#1}}
-
-% typographic conventions
-\newcommand{\qt}[1]{``{#1}''}
-
-% verbatim text
-\newcommand{\isatypewriter}{\fontsize{9pt}{0pt}\tt\renewcommand{\baselinestretch}{1}\setlength{\baselineskip}{9pt}}
-
-% invisibility
 \isadroptag{theory}
 
-% quoted segments
-\makeatletter
-\isakeeptag{quote}
-\newenvironment{quotesegment}{\begin{quote}\isa@parindent\parindent\parindent0pt\isa@parskip\parskip\parskip0pt}{\end{quote}}
-\renewcommand{\isatagquote}{\begin{quotesegment}}
-\renewcommand{\endisatagquote}{\end{quotesegment}}
-\makeatother
-
-%\renewcommand{\isasymlongleftrightarrow}{\isamath{\leftrightarrow}}
-%\renewcommand{\isasymdiv}{\isamath{{}^{-1}}}
-%\renewcommand{\isasymotimes}{\isamath{\circ}}
-
-
-%% content
-
 \title{\includegraphics[scale=0.5]{isabelle_isar}
   \\[4ex] Haskell-style type classes with Isabelle/Isar}
 \author{\emph{Florian Haftmann}}
@@ -69,7 +37,6 @@
 \input{Thy/document/Classes.tex}
 
 \begingroup
-%\tocentry{\bibname}
 \bibliographystyle{plain} \small\raggedright\frenchspacing
 \bibliography{../../manual}
 \endgroup
--- a/doc-src/IsarAdvanced/Classes/style.sty	Wed Dec 10 10:23:47 2008 +0100
+++ b/doc-src/IsarAdvanced/Classes/style.sty	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,3 @@
-
-%% $Id$
 
 %% toc
 \newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1}
@@ -7,54 +5,37 @@
 
 %% references
 \newcommand{\secref}[1]{\S\ref{#1}}
-\newcommand{\chref}[1]{chapter~\ref{#1}}
 \newcommand{\figref}[1]{figure~\ref{#1}}
 
-%% glossary
-\renewcommand{\glossary}[2]{\nomenclature{\bf #1}{#2}}
-\newcommand{\seeglossary}[1]{\emph{#1}}
-\newcommand{\glossaryname}{Glossary}
-\renewcommand{\nomname}{\glossaryname}
-\renewcommand{\pagedeclaration}[1]{\nobreak\quad\dotfill~page~\bold{#1}}
+%% logical markup
+\newcommand{\strong}[1]{{\bfseries {#1}}}
+\newcommand{\qn}[1]{\emph{#1}}
+
+%% typographic conventions
+\newcommand{\qt}[1]{``{#1}''}
+
+%% verbatim text
+\newcommand{\isatypewriter}{\fontsize{9pt}{0pt}\tt\renewcommand{\baselinestretch}{1}\setlength{\baselineskip}{9pt}}
 
-%% index
-\newcommand{\indexml}[1]{\index{\emph{#1}|bold}}
-\newcommand{\indexmltype}[1]{\index{\emph{#1} (type)|bold}}
-\newcommand{\indexmlstructure}[1]{\index{\emph{#1} (structure)|bold}}
-\newcommand{\indexmlfunctor}[1]{\index{\emph{#1} (functor)|bold}}
+%% quoted segments
+\makeatletter
+\isakeeptag{quote}
+\newenvironment{quotesegment}{\begin{quote}\isa@parindent\parindent\parindent0pt\isa@parskip\parskip\parskip0pt}{\end{quote}}
+\renewcommand{\isatagquote}{\begin{quotesegment}}
+\renewcommand{\endisatagquote}{\end{quotesegment}}
+\makeatother
 
-%% math
-\newcommand{\text}[1]{\mbox{#1}}
-\newcommand{\isasymvartheta}{\isamath{\theta}}
-\newcommand{\isactrlvec}[1]{\emph{$\overline{#1}$}}
-
+%% presentation
 \setcounter{secnumdepth}{2} \setcounter{tocdepth}{2}
 
 \pagestyle{headings}
-\sloppy
 \binperiod
 \underscoreon
 
 \renewcommand{\isadigit}[1]{\isamath{#1}}
 
-\newenvironment{mldecls}{\par\noindent\begingroup\footnotesize\def\isanewline{\\}\begin{tabular}{l}}{\end{tabular}\smallskip\endgroup}
-
-\isafoldtag{FIXME}
-\isakeeptag{mlref}
-\renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Reference}\begingroup\def\isastyletext{\rm}\small}
-\renewcommand{\endisatagmlref}{\endgroup}
+\isabellestyle{it}
 
-\newcommand{\isasymGUESS}{\isakeyword{guess}}
-\newcommand{\isasymOBTAIN}{\isakeyword{obtain}}
-\newcommand{\isasymTHEORY}{\isakeyword{theory}}
-\newcommand{\isasymUSES}{\isakeyword{uses}}
-\newcommand{\isasymEND}{\isakeyword{end}}
-\newcommand{\isasymCONSTS}{\isakeyword{consts}}
-\newcommand{\isasymDEFS}{\isakeyword{defs}}
-\newcommand{\isasymTHEOREM}{\isakeyword{theorem}}
-\newcommand{\isasymDEFINITION}{\isakeyword{definition}}
-
-\isabellestyle{it}
 
 %%% Local Variables: 
 %%% mode: latex
--- a/doc-src/IsarAdvanced/Codegen/codegen.tex	Wed Dec 10 10:23:47 2008 +0100
+++ b/doc-src/IsarAdvanced/Codegen/codegen.tex	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,3 @@
-
-%% $Id$
 
 \documentclass[12pt,a4paper,fleqn]{report}
 \usepackage{latexsym,graphicx}
@@ -12,43 +10,10 @@
 \usepackage{tikz}
 \usepackage{../../pdfsetup}
 
-%% setup
-
-% hyphenation
 \hyphenation{Isabelle}
 \hyphenation{Isar}
-
-% logical markup
-\newcommand{\strong}[1]{{\bfseries {#1}}}
-\newcommand{\qn}[1]{\emph{#1}}
-
-% typographic conventions
-\newcommand{\qt}[1]{``{#1}''}
-
-% verbatim text
-\newcommand{\isatypewriter}{\fontsize{9pt}{0pt}\tt\renewcommand{\baselinestretch}{1}\setlength{\baselineskip}{9pt}}
-
-% invisibility
 \isadroptag{theory}
 
-% quoted segments
-\makeatletter
-\isakeeptag{quote}
-\newenvironment{quotesegment}{\begin{quote}\isa@parindent\parindent\parindent0pt\isa@parskip\parskip\parskip0pt}{\end{quote}}
-\renewcommand{\isatagquote}{\begin{quotesegment}}
-\renewcommand{\endisatagquote}{\end{quotesegment}}
-\makeatother
-
-\isakeeptag{quotett}
-\renewcommand{\isatagquotett}{\begin{quotesegment}\isabellestyle{tt}\isastyle}
-\renewcommand{\endisatagquotett}{\end{quotesegment}\isabellestyle{it}\isastyle}
-
-% hack
-\newcommand{\isasymSML}{SML}
-
-
-%% contents
-
 \title{\includegraphics[scale=0.5]{isabelle_isar}
   \\[4ex] Code generation from Isabelle/HOL theories}
 \author{\emph{Florian Haftmann}}
@@ -75,7 +40,6 @@
 \input{Thy/document/ML.tex}
 
 \begingroup
-%\tocentry{\bibname}
 \bibliographystyle{plain} \small\raggedright\frenchspacing
 \bibliography{../../manual}
 \endgroup
--- a/doc-src/IsarAdvanced/Codegen/style.sty	Wed Dec 10 10:23:47 2008 +0100
+++ b/doc-src/IsarAdvanced/Codegen/style.sty	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,3 @@
-
-%% $Id$
 
 %% toc
 \newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1}
@@ -7,15 +5,6 @@
 
 %% references
 \newcommand{\secref}[1]{\S\ref{#1}}
-\newcommand{\chref}[1]{chapter~\ref{#1}}
-\newcommand{\figref}[1]{figure~\ref{#1}}
-
-%% glossary
-\renewcommand{\glossary}[2]{\nomenclature{\bf #1}{#2}}
-\newcommand{\seeglossary}[1]{\emph{#1}}
-\newcommand{\glossaryname}{Glossary}
-\renewcommand{\nomname}{\glossaryname}
-\renewcommand{\pagedeclaration}[1]{\nobreak\quad\dotfill~page~\bold{#1}}
 
 %% index
 \newcommand{\indexml}[1]{\index{\emph{#1}|bold}}
@@ -23,38 +12,49 @@
 \newcommand{\indexmlstructure}[1]{\index{\emph{#1} (structure)|bold}}
 \newcommand{\indexmlfunctor}[1]{\index{\emph{#1} (functor)|bold}}
 
-%% math
-\newcommand{\text}[1]{\mbox{#1}}
-\newcommand{\isasymvartheta}{\isamath{\theta}}
-\newcommand{\isactrlvec}[1]{\emph{$\overline{#1}$}}
+%% logical markup
+\newcommand{\strong}[1]{{\bfseries {#1}}}
+\newcommand{\qn}[1]{\emph{#1}}
+
+%% typographic conventions
+\newcommand{\qt}[1]{``{#1}''}
+
+%% verbatim text
+\newcommand{\isatypewriter}{\fontsize{9pt}{0pt}\tt\renewcommand{\baselinestretch}{1}\setlength{\baselineskip}{9pt}}
 
+%% quoted segments
+\makeatletter
+\isakeeptag{quote}
+\newenvironment{quotesegment}{\begin{quote}\isa@parindent\parindent\parindent0pt\isa@parskip\parskip\parskip0pt}{\end{quote}}
+\renewcommand{\isatagquote}{\begin{quotesegment}}
+\renewcommand{\endisatagquote}{\end{quotesegment}}
+\makeatother
+
+\isakeeptag{quotett}
+\renewcommand{\isatagquotett}{\begin{quotesegment}\isabellestyle{tt}\isastyle}
+\renewcommand{\endisatagquotett}{\end{quotesegment}\isabellestyle{it}\isastyle}
+
+%% a trick
+\newcommand{\isasymSML}{SML}
+
+%% presentation
 \setcounter{secnumdepth}{2} \setcounter{tocdepth}{2}
 
 \pagestyle{headings}
-\sloppy
 \binperiod
 \underscoreon
 
 \renewcommand{\isadigit}[1]{\isamath{#1}}
 
+%% ml reference
 \newenvironment{mldecls}{\par\noindent\begingroup\footnotesize\def\isanewline{\\}\begin{tabular}{l}}{\end{tabular}\smallskip\endgroup}
 
-\isafoldtag{FIXME}
 \isakeeptag{mlref}
 \renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Reference}\begingroup\def\isastyletext{\rm}\small}
 \renewcommand{\endisatagmlref}{\endgroup}
 
-\newcommand{\isasymGUESS}{\isakeyword{guess}}
-\newcommand{\isasymOBTAIN}{\isakeyword{obtain}}
-\newcommand{\isasymTHEORY}{\isakeyword{theory}}
-\newcommand{\isasymUSES}{\isakeyword{uses}}
-\newcommand{\isasymEND}{\isakeyword{end}}
-\newcommand{\isasymCONSTS}{\isakeyword{consts}}
-\newcommand{\isasymDEFS}{\isakeyword{defs}}
-\newcommand{\isasymTHEOREM}{\isakeyword{theorem}}
-\newcommand{\isasymDEFINITION}{\isakeyword{definition}}
+\isabellestyle{it}
 
-\isabellestyle{it}
 
 %%% Local Variables: 
 %%% mode: latex
--- a/doc-src/IsarImplementation/Thy/ML.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/doc-src/IsarImplementation/Thy/ML.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -306,7 +306,7 @@
   a theory by constant declararion and primitive definitions:
 
   \smallskip\begin{mldecls}
-  @{ML "Sign.declare_const: Properties.T -> (Name.binding * typ) * mixfix
+  @{ML "Sign.declare_const: Properties.T -> (Binding.T * typ) * mixfix
   -> theory -> term * theory"} \\
   @{ML "Thm.add_def: bool -> bool -> bstring * term -> theory -> thm * theory"}
   \end{mldecls}
@@ -319,7 +319,7 @@
   @{ML "(fn (t, thy) => Thm.add_def false false
   (\"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})) thy)
     (Sign.declare_const []
-      ((Name.binding \"bar\", @{typ \"foo => foo\"}), NoSyn) thy)"}
+      ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn) thy)"}
   \end{mldecls}
 
   \noindent With increasing numbers of applications, this code gets quite
@@ -333,7 +333,7 @@
 
   \smallskip\begin{mldecls}
 @{ML "thy
-|> Sign.declare_const [] ((Name.binding \"bar\", @{typ \"foo => foo\"}), NoSyn)
+|> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
 |> (fn (t, thy) => thy
 |> Thm.add_def false false
      (\"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))"}
@@ -357,7 +357,7 @@
 
   \smallskip\begin{mldecls}
 @{ML "thy
-|> Sign.declare_const [] ((Name.binding \"bar\", @{typ \"foo => foo\"}), NoSyn)
+|> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
 |-> (fn t => Thm.add_def false false
       (\"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))
 "}
@@ -367,7 +367,7 @@
 
   \smallskip\begin{mldecls}
 @{ML "thy
-|> Sign.declare_const [] ((Name.binding \"bar\", @{typ \"foo => foo\"}), NoSyn)
+|> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
 |>> (fn t => Logic.mk_equals (t, @{term \"%x. x\"}))
 |-> (fn def => Thm.add_def false false (\"bar_def\", def))
 "}
@@ -378,7 +378,7 @@
 
   \smallskip\begin{mldecls}
 @{ML "thy
-|> Sign.declare_const [] ((Name.binding \"bar\", @{typ \"foo => foo\"}), NoSyn)
+|> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
 ||> Sign.add_path \"foobar\"
 |-> (fn t => Thm.add_def false false
       (\"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))
@@ -390,8 +390,8 @@
 
   \smallskip\begin{mldecls}
 @{ML "thy
-|> Sign.declare_const [] ((Name.binding \"bar\", @{typ \"foo => foo\"}), NoSyn)
-||>> Sign.declare_const [] ((Name.binding \"foobar\", @{typ \"foo => foo\"}), NoSyn)
+|> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
+||>> Sign.declare_const [] ((Binding.name \"foobar\", @{typ \"foo => foo\"}), NoSyn)
 |-> (fn (t1, t2) => Thm.add_def false false
       (\"bar_def\", Logic.mk_equals (t1, t2)))
 "}
@@ -437,7 +437,7 @@
 in
   thy
   |> fold_map (fn const => Sign.declare_const []
-       ((Name.binding const, @{typ \"foo => foo\"}), NoSyn)) consts
+       ((Binding.name const, @{typ \"foo => foo\"}), NoSyn)) consts
   |>> map (fn t => Logic.mk_equals (t, @{term \"%x. x\"}))
   |-> (fn defs => fold_map (fn def =>
        Thm.add_def false false (\"\", def)) defs)
@@ -475,12 +475,12 @@
   \smallskip\begin{mldecls}
 @{ML "thy
 |> tap (fn _ => writeln \"now adding constant\")
-|> Sign.declare_const [] ((Name.binding \"bar\", @{typ \"foo => foo\"}), NoSyn)
+|> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
 ||>> `(fn thy => Sign.declared_const thy
-         (Sign.full_name thy \"foobar\"))
+         (Sign.full_name thy (Binding.name \"foobar\")))
 |-> (fn (t, b) => if b then I
        else Sign.declare_const []
-         ((Name.binding \"foobar\", @{typ \"foo => foo\"}), NoSyn) #> snd)
+         ((Binding.name \"foobar\", @{typ \"foo => foo\"}), NoSyn) #> snd)
 "}
   \end{mldecls}
 *}
--- a/doc-src/IsarImplementation/Thy/document/ML.tex	Wed Dec 10 10:23:47 2008 +0100
+++ b/doc-src/IsarImplementation/Thy/document/ML.tex	Wed Dec 10 10:28:16 2008 +0100
@@ -358,7 +358,7 @@
   a theory by constant declararion and primitive definitions:
 
   \smallskip\begin{mldecls}
-  \verb|Sign.declare_const: Properties.T -> (Name.binding * typ) * mixfix|\isasep\isanewline%
+  \verb|Sign.declare_const: Properties.T -> (Binding.T * typ) * mixfix|\isasep\isanewline%
 \verb|  -> theory -> term * theory| \\
   \verb|Thm.add_def: bool -> bool -> bstring * term -> theory -> thm * theory|
   \end{mldecls}
@@ -371,7 +371,7 @@
   \verb|(fn (t, thy) => Thm.add_def false false|\isasep\isanewline%
 \verb|  ("bar_def", Logic.mk_equals (t, @{term "%x. x"})) thy)|\isasep\isanewline%
 \verb|    (Sign.declare_const []|\isasep\isanewline%
-\verb|      ((Name.binding "bar", @{typ "foo => foo"}), NoSyn) thy)|
+\verb|      ((Binding.name "bar", @{typ "foo => foo"}), NoSyn) thy)|
   \end{mldecls}
 
   \noindent With increasing numbers of applications, this code gets quite
@@ -386,7 +386,7 @@
 
   \smallskip\begin{mldecls}
 \verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const [] ((Name.binding "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
+\verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb|> (fn (t, thy) => thy|\isasep\isanewline%
 \verb||\verb,|,\verb|> Thm.add_def false false|\isasep\isanewline%
 \verb|     ("bar_def", Logic.mk_equals (t, @{term "%x. x"})))|
@@ -425,7 +425,7 @@
 
   \smallskip\begin{mldecls}
 \verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const [] ((Name.binding "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
+\verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb|-> (fn t => Thm.add_def false false|\isasep\isanewline%
 \verb|      ("bar_def", Logic.mk_equals (t, @{term "%x. x"})))|\isasep\isanewline%
 
@@ -435,7 +435,7 @@
 
   \smallskip\begin{mldecls}
 \verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const [] ((Name.binding "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
+\verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb|>> (fn t => Logic.mk_equals (t, @{term "%x. x"}))|\isasep\isanewline%
 \verb||\verb,|,\verb|-> (fn def => Thm.add_def false false ("bar_def", def))|\isasep\isanewline%
 
@@ -446,7 +446,7 @@
 
   \smallskip\begin{mldecls}
 \verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const [] ((Name.binding "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
+\verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb||\verb,|,\verb|> Sign.add_path "foobar"|\isasep\isanewline%
 \verb||\verb,|,\verb|-> (fn t => Thm.add_def false false|\isasep\isanewline%
 \verb|      ("bar_def", Logic.mk_equals (t, @{term "%x. x"})))|\isasep\isanewline%
@@ -458,8 +458,8 @@
 
   \smallskip\begin{mldecls}
 \verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const [] ((Name.binding "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
-\verb||\verb,|,\verb||\verb,|,\verb|>> Sign.declare_const [] ((Name.binding "foobar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
+\verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
+\verb||\verb,|,\verb||\verb,|,\verb|>> Sign.declare_const [] ((Binding.name "foobar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb|-> (fn (t1, t2) => Thm.add_def false false|\isasep\isanewline%
 \verb|      ("bar_def", Logic.mk_equals (t1, t2)))|\isasep\isanewline%
 
@@ -520,7 +520,7 @@
 \verb|in|\isasep\isanewline%
 \verb|  thy|\isasep\isanewline%
 \verb|  |\verb,|,\verb|> fold_map (fn const => Sign.declare_const []|\isasep\isanewline%
-\verb|       ((Name.binding const, @{typ "foo => foo"}), NoSyn)) consts|\isasep\isanewline%
+\verb|       ((Binding.name const, @{typ "foo => foo"}), NoSyn)) consts|\isasep\isanewline%
 \verb|  |\verb,|,\verb|>> map (fn t => Logic.mk_equals (t, @{term "%x. x"}))|\isasep\isanewline%
 \verb|  |\verb,|,\verb|-> (fn defs => fold_map (fn def =>|\isasep\isanewline%
 \verb|       Thm.add_def false false ("", def)) defs)|\isasep\isanewline%
@@ -588,12 +588,12 @@
   \smallskip\begin{mldecls}
 \verb|thy|\isasep\isanewline%
 \verb||\verb,|,\verb|> tap (fn _ => writeln "now adding constant")|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const [] ((Name.binding "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
+\verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb||\verb,|,\verb|>> `(fn thy => Sign.declared_const thy|\isasep\isanewline%
-\verb|         (Sign.full_name thy "foobar"))|\isasep\isanewline%
+\verb|         (Sign.full_name thy (Binding.name "foobar")))|\isasep\isanewline%
 \verb||\verb,|,\verb|-> (fn (t, b) => if b then I|\isasep\isanewline%
 \verb|       else Sign.declare_const []|\isasep\isanewline%
-\verb|         ((Name.binding "foobar", @{typ "foo => foo"}), NoSyn) #> snd)|\isasep\isanewline%
+\verb|         ((Binding.name "foobar", @{typ "foo => foo"}), NoSyn) #> snd)|\isasep\isanewline%
 
   \end{mldecls}%
 \end{isamarkuptext}%
--- a/doc-src/IsarImplementation/Thy/document/logic.tex	Wed Dec 10 10:23:47 2008 +0100
+++ b/doc-src/IsarImplementation/Thy/document/logic.tex	Wed Dec 10 10:28:16 2008 +0100
@@ -3,9 +3,6 @@
 \def\isabellecontext{logic}%
 %
 \isadelimtheory
-\isanewline
-\isanewline
-\isanewline
 %
 \endisadelimtheory
 %
@@ -328,9 +325,9 @@
   \indexml{fastype\_of}\verb|fastype_of: term -> typ| \\
   \indexml{lambda}\verb|lambda: term -> term -> term| \\
   \indexml{betapply}\verb|betapply: term * term -> term| \\
-  \indexml{Sign.declare\_const}\verb|Sign.declare_const: Properties.T -> (Name.binding * typ) * mixfix ->|\isasep\isanewline%
+  \indexml{Sign.declare\_const}\verb|Sign.declare_const: Properties.T -> (Binding.T * typ) * mixfix ->|\isasep\isanewline%
 \verb|  theory -> term * theory| \\
-  \indexml{Sign.add\_abbrev}\verb|Sign.add_abbrev: string -> Properties.T -> Name.binding * term ->|\isasep\isanewline%
+  \indexml{Sign.add\_abbrev}\verb|Sign.add_abbrev: string -> Properties.T -> Binding.T * term ->|\isasep\isanewline%
 \verb|  theory -> (term * term) * theory| \\
   \indexml{Sign.const\_typargs}\verb|Sign.const_typargs: theory -> string * typ -> typ list| \\
   \indexml{Sign.const\_instance}\verb|Sign.const_instance: theory -> string * typ list -> typ| \\
--- a/doc-src/IsarImplementation/Thy/document/prelim.tex	Wed Dec 10 10:23:47 2008 +0100
+++ b/doc-src/IsarImplementation/Thy/document/prelim.tex	Wed Dec 10 10:28:16 2008 +0100
@@ -816,13 +816,13 @@
   \indexmltype{NameSpace.naming}\verb|type NameSpace.naming| \\
   \indexml{NameSpace.default\_naming}\verb|NameSpace.default_naming: NameSpace.naming| \\
   \indexml{NameSpace.add\_path}\verb|NameSpace.add_path: string -> NameSpace.naming -> NameSpace.naming| \\
-  \indexml{NameSpace.full}\verb|NameSpace.full: NameSpace.naming -> string -> string| \\
+  \indexml{NameSpace.full\_name}\verb|NameSpace.full_name: NameSpace.naming -> Binding.T -> string| \\
   \end{mldecls}
   \begin{mldecls}
   \indexmltype{NameSpace.T}\verb|type NameSpace.T| \\
   \indexml{NameSpace.empty}\verb|NameSpace.empty: NameSpace.T| \\
   \indexml{NameSpace.merge}\verb|NameSpace.merge: NameSpace.T * NameSpace.T -> NameSpace.T| \\
-  \indexml{NameSpace.declare}\verb|NameSpace.declare: NameSpace.naming -> string -> NameSpace.T -> NameSpace.T| \\
+  \indexml{NameSpace.declare}\verb|NameSpace.declare: NameSpace.naming -> Binding.T -> NameSpace.T -> string * NameSpace.T| \\
   \indexml{NameSpace.intern}\verb|NameSpace.intern: NameSpace.T -> string -> string| \\
   \indexml{NameSpace.extern}\verb|NameSpace.extern: NameSpace.T -> string -> string| \\
   \end{mldecls}
@@ -851,9 +851,9 @@
   \item \verb|NameSpace.add_path|~\isa{path\ naming} augments the
   naming policy by extending its path component.
 
-  \item \verb|NameSpace.full|\isa{naming\ name} turns a name
-  specification (usually a basic name) into the fully qualified
-  internal version, according to the given naming policy.
+  \item \verb|NameSpace.full_name|\isa{naming\ binding} turns a name
+  binding (usually a basic name) into the fully qualified
+  internal name, according to the given naming policy.
 
   \item \verb|NameSpace.T| represents name spaces.
 
@@ -861,15 +861,16 @@
   maintaining name spaces according to theory data management
   (\secref{sec:context-data}).
 
-  \item \verb|NameSpace.declare|~\isa{naming\ name\ space} enters a
-  fully qualified name into the name space, with external accesses
-  determined by the naming policy.
+  \item \verb|NameSpace.declare|~\isa{naming\ bindings\ space} enters a
+  name binding as fully qualified internal name into the name space,
+  with external accesses determined by the naming policy.
 
   \item \verb|NameSpace.intern|~\isa{space\ name} internalizes a
   (partially qualified) external name.
 
   This operation is mostly for parsing!  Note that fully qualified
-  names stemming from declarations are produced via \verb|NameSpace.full| (or its derivatives for \verb|theory| and
+  names stemming from declarations are produced via \verb|NameSpace.full_name| and \verb|NameSpace.declare|
+  (or their derivatives for \verb|theory| and
   \verb|Proof.context|).
 
   \item \verb|NameSpace.extern|~\isa{space\ name} externalizes a
--- a/doc-src/IsarImplementation/Thy/logic.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/doc-src/IsarImplementation/Thy/logic.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -1,6 +1,3 @@
-
-(* $Id$ *)
-
 theory logic imports base begin
 
 chapter {* Primitive logic \label{ch:logic} *}
@@ -326,9 +323,9 @@
   @{index_ML fastype_of: "term -> typ"} \\
   @{index_ML lambda: "term -> term -> term"} \\
   @{index_ML betapply: "term * term -> term"} \\
-  @{index_ML Sign.declare_const: "Properties.T -> (Name.binding * typ) * mixfix ->
+  @{index_ML Sign.declare_const: "Properties.T -> (Binding.T * typ) * mixfix ->
   theory -> term * theory"} \\
-  @{index_ML Sign.add_abbrev: "string -> Properties.T -> Name.binding * term ->
+  @{index_ML Sign.add_abbrev: "string -> Properties.T -> Binding.T * term ->
   theory -> (term * term) * theory"} \\
   @{index_ML Sign.const_typargs: "theory -> string * typ -> typ list"} \\
   @{index_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\
--- a/doc-src/IsarImplementation/Thy/prelim.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/doc-src/IsarImplementation/Thy/prelim.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -707,13 +707,13 @@
   @{index_ML_type NameSpace.naming} \\
   @{index_ML NameSpace.default_naming: NameSpace.naming} \\
   @{index_ML NameSpace.add_path: "string -> NameSpace.naming -> NameSpace.naming"} \\
-  @{index_ML NameSpace.full: "NameSpace.naming -> string -> string"} \\
+  @{index_ML NameSpace.full_name: "NameSpace.naming -> Binding.T -> string"} \\
   \end{mldecls}
   \begin{mldecls}
   @{index_ML_type NameSpace.T} \\
   @{index_ML NameSpace.empty: NameSpace.T} \\
   @{index_ML NameSpace.merge: "NameSpace.T * NameSpace.T -> NameSpace.T"} \\
-  @{index_ML NameSpace.declare: "NameSpace.naming -> string -> NameSpace.T -> NameSpace.T"} \\
+  @{index_ML NameSpace.declare: "NameSpace.naming -> Binding.T -> NameSpace.T -> string * NameSpace.T"} \\
   @{index_ML NameSpace.intern: "NameSpace.T -> string -> string"} \\
   @{index_ML NameSpace.extern: "NameSpace.T -> string -> string"} \\
   \end{mldecls}
@@ -743,9 +743,9 @@
   \item @{ML NameSpace.add_path}~@{text "path naming"} augments the
   naming policy by extending its path component.
 
-  \item @{ML NameSpace.full}@{text "naming name"} turns a name
-  specification (usually a basic name) into the fully qualified
-  internal version, according to the given naming policy.
+  \item @{ML NameSpace.full_name}@{text "naming binding"} turns a name
+  binding (usually a basic name) into the fully qualified
+  internal name, according to the given naming policy.
 
   \item @{ML_type NameSpace.T} represents name spaces.
 
@@ -754,16 +754,17 @@
   maintaining name spaces according to theory data management
   (\secref{sec:context-data}).
 
-  \item @{ML NameSpace.declare}~@{text "naming name space"} enters a
-  fully qualified name into the name space, with external accesses
-  determined by the naming policy.
+  \item @{ML NameSpace.declare}~@{text "naming bindings space"} enters a
+  name binding as fully qualified internal name into the name space,
+  with external accesses determined by the naming policy.
 
   \item @{ML NameSpace.intern}~@{text "space name"} internalizes a
   (partially qualified) external name.
 
   This operation is mostly for parsing!  Note that fully qualified
   names stemming from declarations are produced via @{ML
-  "NameSpace.full"} (or its derivatives for @{ML_type theory} and
+  "NameSpace.full_name"} and @{ML "NameSpace.declare"}
+  (or their derivatives for @{ML_type theory} and
   @{ML_type Proof.context}).
 
   \item @{ML NameSpace.extern}~@{text "space name"} externalizes a
--- a/etc/settings	Wed Dec 10 10:23:47 2008 +0100
+++ b/etc/settings	Wed Dec 10 10:28:16 2008 +0100
@@ -75,8 +75,11 @@
 ISABELLE_SCALA="scala"
 ISABELLE_JAVA="java"
 
-[ -e "$ISABELLE_HOME/contrib/scala" ] && \
+if [ -e "$ISABELLE_HOME/contrib/scala" ]; then
   classpath "$ISABELLE_HOME/contrib/scala/lib/scala-library.jar"
+elif [ -e "$ISABELLE_HOME/../scala" ]; then
+  classpath "$ISABELLE_HOME/../scala/lib/scala-library.jar"
+fi
 
 classpath "$ISABELLE_HOME/lib/classes/Pure.jar"
 
@@ -232,10 +235,22 @@
 ## Set HOME only for tools you have installed!
 
 # External provers
-E_HOME=$(choosefrom "$ISABELLE_HOME/contrib/E/$ML_PLATFORM" "/usr/local/E" "")
-VAMPIRE_HOME=$(choosefrom "$ISABELLE_HOME/contrib/vampire/$ML_PLATFORM" "/usr/local/Vampire" \
-  "$ISABELLE_HOME/contrib/SystemOnTPTP" "")
-SPASS_HOME=$(choosefrom "$ISABELLE_HOME/contrib/spass/$ML_PLATFORM/bin" "/usr/local/SPASS" "")
+E_HOME=$(choosefrom \
+  "$ISABELLE_HOME/contrib/E/$ML_PLATFORM" \
+  "$ISABELLE_HOME/../E/$ML_PLATFORM" \
+  "/usr/local/E" \
+  "")
+VAMPIRE_HOME=$(choosefrom \
+  "$ISABELLE_HOME/contrib/vampire/$ML_PLATFORM" \
+  "$ISABELLE_HOME/../vampire/$ML_PLATFORM" \
+  "/usr/local/Vampire" \
+  "$ISABELLE_HOME/contrib/SystemOnTPTP" \
+  "")
+SPASS_HOME=$(choosefrom \
+  "$ISABELLE_HOME/contrib/spass/$ML_PLATFORM/bin" \
+  "$ISABELLE_HOME/../spass/$ML_PLATFORM/bin" \
+  "/usr/local/SPASS" \
+  "")
 
 # HOL4 proof objects (cf. Isabelle/src/HOL/Import)
 #HOL4_PROOFS="$ISABELLE_HOME_USER/proofs:$ISABELLE_HOME/proofs"
--- a/src/FOL/ex/NewLocaleSetup.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/FOL/ex/NewLocaleSetup.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,4 @@
 (*  Title:      FOL/ex/NewLocaleSetup.thy
-    ID:         $Id$
     Author:     Clemens Ballarin, TU Muenchen
 
 Testing environment for locale expressions --- experimental.
@@ -40,6 +39,21 @@
    Toplevel.unknown_theory o Toplevel.keep (fn state =>
      NewLocale.print_locale (Toplevel.theory_of state) show_facts name))));
 
+val _ =
+  OuterSyntax.command "interpretation"
+    "prove interpretation of locale expression in theory" K.thy_goal
+    (P.!!! Expression.parse_expression
+        >> (fn expr => Toplevel.print o
+            Toplevel.theory_to_proof (Expression.interpretation_cmd expr)));
+
+val _ =
+  OuterSyntax.command "interpret"
+    "prove interpretation of locale expression in proof context"
+    (K.tag_proof K.prf_goal)
+    (P.!!! Expression.parse_expression
+        >> (fn expr => Toplevel.print o
+            Toplevel.proof' (fn int => Expression.interpret_cmd expr int)));
+
 end
 
 *}
--- a/src/FOL/ex/NewLocaleTest.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/FOL/ex/NewLocaleTest.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,4 @@
 (*  Title:      FOL/ex/NewLocaleTest.thy
-    ID:         $Id$
     Author:     Clemens Ballarin, TU Muenchen
 
 Testing environment for locale expressions --- experimental.
@@ -19,6 +18,11 @@
   zero :: int ("0")
   minus :: "int => int" ("- _")
 
+axioms
+  int_assoc: "(x + y::int) + z = x + (y + z)"
+  int_zero: "0 + x = x"
+  int_minus: "(-x) + x = 0"
+  int_minus2: "-(-x) = x"
 
 text {* Inference of parameter types *}
 
@@ -30,7 +34,7 @@
 
 (*
 locale param_top = param2 r for r :: "'b :: {}"
-print_locale! param_top
+  Fails, cannot generalise parameter.
 *)
 
 locale param3 = fixes p (infix ".." 50)
@@ -109,6 +113,23 @@
 print_locale! use_decl thm use_decl_def
 
 
+text {* Defines *}
+
+locale logic_def =
+  fixes land (infixl "&&" 55)
+    and lor (infixl "||" 50)
+    and lnot ("-- _" [60] 60)
+  assumes assoc: "(x && y) && z = x && (y && z)"
+    and notnot: "-- (-- x) = x"
+  defines "x || y == --(-- x && -- y)"
+begin
+
+lemma "x || y = --(-- x && --y)"
+  by (unfold lor_def) (rule refl)
+
+end
+
+
 text {* Theorem statements *}
 
 lemma (in lgrp) lcancel:
@@ -298,4 +319,63 @@
 
 print_locale! trivial  (* No instance for y created (subsumed). *)
 
+
+text {* Sublocale, then interpretation in theory *}
+
+interpretation int: lgrp "op +" "0" "minus"
+proof unfold_locales
+qed (rule int_assoc int_zero int_minus)+
+
+thm int.assoc int.semi_axioms
+
+interpretation int2: semi "op +"
+  by unfold_locales  (* subsumed, thm int2.assoc not generated *)
+
+thm int.lone int.right.rone
+  (* the latter comes through the sublocale relation *)
+
+
+text {* Interpretation in theory, then sublocale *}
+
+interpretation (* fol: *) logic "op +" "minus"
+(* FIXME declaration of qualified names *)
+  by unfold_locales (rule int_assoc int_minus2)+
+
+locale logic2 =
+  fixes land (infixl "&&" 55)
+    and lnot ("-- _" [60] 60)
+  assumes assoc: "(x && y) && z = x && (y && z)"
+    and notnot: "-- (-- x) = x"
+begin
+(* FIXME
+definition lor (infixl "||" 50) where
+  "x || y = --(-- x && -- y)"
+*)
 end
+
+sublocale logic < two: logic2
+  by unfold_locales (rule assoc notnot)+
+
+thm two.assoc
+
+
+text {* Interpretation in proofs *}
+
+lemma True
+proof
+  interpret "local": lgrp "op +" "0" "minus"
+    by unfold_locales  (* subsumed *)
+  {
+    fix zero :: int
+    assume "!!x. zero + x = x" "!!x. (-x) + x = zero"
+    then interpret local_fixed: lgrp "op +" zero "minus"
+      by unfold_locales
+    thm local_fixed.lone
+  }
+  assume "!!x zero. zero + x = x" "!!x zero. (-x) + x = zero"
+  then interpret local_free: lgrp "op +" zero "minus" for zero
+    by unfold_locales
+  thm local_free.lone [where ?zero = 0]
+qed
+
+end
--- a/src/HOL/Arith_Tools.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Arith_Tools.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -40,7 +40,7 @@
 text{*No longer required as a simprule because of the @{text inverse_fold}
    simproc*}
 lemma Suc_diff_number_of:
-     "neg (number_of (uminus v)::int) ==>
+     "Int.Pls < v ==>
       Suc m - (number_of v) = m - (number_of (Int.pred v))"
 apply (subst Suc_diff_eq_diff_pred)
 apply simp
--- a/src/HOL/Code_Eval.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Code_Eval.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -63,7 +63,7 @@
       thy
       |> TheoryTarget.instantiation ([tyco], vs, @{sort term_of})
       |> `(fn lthy => Syntax.check_term lthy eq)
-      |-> (fn eq => Specification.definition (NONE, ((Name.binding (triv_name_of eq), []), eq)))
+      |-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq)))
       |> snd
       |> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
       |> LocalTheory.exit_global
--- a/src/HOL/Datatype.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Datatype.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -605,6 +605,9 @@
 lemma insert_None_conv_UNIV: "insert None (range Some) = UNIV"
   by (rule set_ext, case_tac x) auto
 
+lemma inj_Some [simp]: "inj_on Some A"
+  by (rule inj_onI) simp
+
 
 subsubsection {* Operations *}
 
--- a/src/HOL/Finite_Set.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Finite_Set.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -1828,6 +1828,18 @@
 by (simp add: card_cartesian_product)
 
 
+subsubsection {* Cardinality of sums *}
+
+lemma card_Plus:
+  assumes "finite A" and "finite B"
+  shows "card (A <+> B) = card A + card B"
+proof -
+  have "Inl`A \<inter> Inr`B = {}" by fast
+  with assms show ?thesis
+    unfolding Plus_def
+    by (simp add: card_Un_disjoint card_image)
+qed
+
 
 subsubsection {* Cardinality of the Powerset *}
 
--- a/src/HOL/Groebner_Basis.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Groebner_Basis.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -169,16 +169,20 @@
   proof qed (auto simp add: ring_simps power_Suc)
 
 lemmas nat_arith =
-  add_nat_number_of diff_nat_number_of mult_nat_number_of eq_nat_number_of less_nat_number_of
+  add_nat_number_of
+  diff_nat_number_of
+  mult_nat_number_of
+  eq_nat_number_of
+  less_nat_number_of
 
 lemma not_iszero_Numeral1: "\<not> iszero (Numeral1::'a::number_ring)"
   by (simp add: numeral_1_eq_1)
+
 lemmas comp_arith = Let_def arith_simps nat_arith rel_simps if_False
   if_True add_0 add_Suc add_number_of_left mult_number_of_left
   numeral_1_eq_1[symmetric] Suc_eq_add_numeral_1
-  numeral_0_eq_0[symmetric] numerals[symmetric] not_iszero_1
-  iszero_number_of_Bit1 iszero_number_of_Bit0 nonzero_number_of_Min
-  iszero_number_of_Pls iszero_0 not_iszero_Numeral1
+  numeral_0_eq_0[symmetric] numerals[symmetric]
+  iszero_simps not_iszero_Numeral1
 
 lemmas semiring_norm = comp_arith
 
@@ -458,7 +462,6 @@
 declare zmod_eq_dvd_iff[algebra]
 declare nat_mod_eq_iff[algebra]
 
-
 subsection{* Groebner Bases for fields *}
 
 interpretation class_fieldgb:
@@ -607,14 +610,6 @@
              @{cpat "((?a::(?'a::{field, ord}))/ ?b) = ?c"}],
              name = "ord_frac_simproc", proc = proc3, identifier = []}
 
-val nat_arith = map thm ["add_nat_number_of", "diff_nat_number_of",
-               "mult_nat_number_of", "eq_nat_number_of", "less_nat_number_of"]
-
-val comp_arith = (map thm ["Let_def", "if_False", "if_True", "add_0",
-                 "add_Suc", "add_number_of_left", "mult_number_of_left",
-                 "Suc_eq_add_numeral_1"])@
-                 (map (fn s => thm s RS sym) ["numeral_1_eq_1", "numeral_0_eq_0"])
-                 @ @{thms arith_simps} @ nat_arith @ @{thms rel_simps}
 val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"},
            @{thm "divide_Numeral1"},
            @{thm "Ring_and_Field.divide_zero"}, @{thm "divide_Numeral0"},
@@ -630,7 +625,7 @@
 in
 val comp_conv = (Simplifier.rewrite
 (HOL_basic_ss addsimps @{thms "Groebner_Basis.comp_arith"}
-              addsimps ths addsimps comp_arith addsimps simp_thms
+              addsimps ths addsimps simp_thms
               addsimprocs field_cancel_numeral_factors
                addsimprocs [add_frac_frac_simproc, add_frac_num_simproc,
                             ord_frac_simproc]
--- a/src/HOL/Import/hol4rews.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Import/hol4rews.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -282,7 +282,7 @@
 	val thy = case opt_get_output_thy thy of
 		      "" => thy
 		    | output_thy => if internal
-				    then add_hol4_cmove (Sign.full_name thy bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
+				    then add_hol4_cmove (Sign.full_bname thy bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
 				    else thy
 	val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
 	val curmaps = HOL4ConstMaps.get thy
@@ -324,7 +324,7 @@
 	val thy = case opt_get_output_thy thy of
 		      "" => thy
 		    | output_thy => if internal
-				    then add_hol4_cmove (Sign.full_name thy bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
+				    then add_hol4_cmove (Sign.full_bname thy bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
 				    else thy
 	val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
 	val curmaps = HOL4ConstMaps.get thy
@@ -348,7 +348,7 @@
 
 fun add_hol4_pending bthy bthm hth thy =
     let
-	val thmname = Sign.full_name thy bthm
+	val thmname = Sign.full_bname thy bthm
 	val _ = message ("Add pending " ^ bthy ^ "." ^ bthm)
 	val curpend = HOL4Pending.get thy
 	val newpend = StringPair.update_new ((bthy,bthm),hth) curpend
--- a/src/HOL/Import/lazy_seq.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Import/lazy_seq.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -80,7 +80,7 @@
 structure LazySeq :> LAZY_SEQ =
 struct
 
-datatype 'a seq = Seq of ('a * 'a seq) option Lazy.T
+datatype 'a seq = Seq of ('a * 'a seq) option lazy
 
 exception Empty
 
--- a/src/HOL/Import/proof_kernel.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Import/proof_kernel.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -1960,7 +1960,7 @@
         val hth = case Shuffler.set_prop thy22 (HOLogic.mk_Trueprop tm24) [("",th)] of
                       SOME (_,res) => HOLThm(rens_of linfo,res)
                     | NONE => raise ERR "new_definition" "Bad conclusion"
-        val fullname = Sign.full_name thy22 thmname
+        val fullname = Sign.full_bname thy22 thmname
         val thy22' = case opt_get_output_thy thy22 of
                          "" => (ImportRecorder.add_hol_mapping thyname thmname fullname;
                                 add_hol4_mapping thyname thmname fullname thy22)
--- a/src/HOL/Int.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Int.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -855,7 +855,7 @@
     add: compare_rls of_nat_1 [symmetric] of_nat_add [symmetric])
 qed
 
-lemma neg_simps:
+lemma bin_less_0_simps:
   "Pls < 0 \<longleftrightarrow> False"
   "Min < 0 \<longleftrightarrow> True"
   "Bit0 w < 0 \<longleftrightarrow> w < 0"
@@ -908,7 +908,7 @@
     less_bin_lemma [of "Bit1 k"]
     less_bin_lemma [of "pred Pls"]
     less_bin_lemma [of "pred k"]
-  by (simp_all add: neg_simps succ_pred)
+  by (simp_all add: bin_less_0_simps succ_pred)
 
 text {* Less-than-or-equal *}
 
@@ -1187,6 +1187,12 @@
     by (auto simp add: iszero_def number_of_eq numeral_simps)
 qed
 
+lemmas iszero_simps =
+  iszero_0 not_iszero_1
+  iszero_number_of_Pls nonzero_number_of_Min
+  iszero_number_of_Bit0 iszero_number_of_Bit1
+(* iszero_number_of_Pls would never normally be used
+   because its lhs simplifies to "iszero 0" *)
 
 subsubsection {* The Less-Than Relation *}
 
@@ -1248,6 +1254,10 @@
   by (simp add: neg_def number_of_eq numeral_simps)
 qed
 
+lemmas neg_simps =
+  not_neg_0 not_neg_1
+  not_neg_number_of_Pls neg_number_of_Min
+  neg_number_of_Bit0 neg_number_of_Bit1
 
 text {* Less-Than or Equals *}
 
@@ -1307,15 +1317,15 @@
   "(number_of x::'a::{ordered_idom,number_ring}) \<le> number_of y \<longleftrightarrow> x \<le> y"
   unfolding number_of_eq by (rule of_int_le_iff)
 
+lemma eq_number_of [simp]:
+  "(number_of x::'a::{ring_char_0,number_ring}) = number_of y \<longleftrightarrow> x = y"
+  unfolding number_of_eq by (rule of_int_eq_iff)
+
 lemmas rel_simps [simp] = 
   less_number_of less_bin_simps
   le_number_of le_bin_simps
-  eq_number_of_eq iszero_0 nonzero_number_of_Min
-  iszero_number_of_Bit0 iszero_number_of_Bit1
-  not_neg_number_of_Pls not_neg_0 not_neg_1 not_iszero_1
-  neg_number_of_Min neg_number_of_Bit0 neg_number_of_Bit1
-(* iszero_number_of_Pls would never be used
-   because its lhs simplifies to "iszero 0" *)
+  eq_number_of_eq eq_bin_simps
+  iszero_simps neg_simps
 
 
 subsubsection {* Simplification of arithmetic when nested to the right. *}
@@ -1576,17 +1586,17 @@
 
 text{*Allow 0 or 1 on either side with a binary numeral on the other*}
 lemmas less_special =
-  binop_eq [of "op <", OF less_number_of_eq_neg numeral_0_eq_0 refl, standard]
-  binop_eq [of "op <", OF less_number_of_eq_neg numeral_1_eq_1 refl, standard]
-  binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_0_eq_0, standard]
-  binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_1_eq_1, standard]
+  binop_eq [of "op <", OF less_number_of numeral_0_eq_0 refl, standard]
+  binop_eq [of "op <", OF less_number_of numeral_1_eq_1 refl, standard]
+  binop_eq [of "op <", OF less_number_of refl numeral_0_eq_0, standard]
+  binop_eq [of "op <", OF less_number_of refl numeral_1_eq_1, standard]
 
 text{*Allow 0 or 1 on either side with a binary numeral on the other*}
 lemmas le_special =
-    binop_eq [of "op \<le>", OF le_number_of_eq numeral_0_eq_0 refl, standard]
-    binop_eq [of "op \<le>", OF le_number_of_eq numeral_1_eq_1 refl, standard]
-    binop_eq [of "op \<le>", OF le_number_of_eq refl numeral_0_eq_0, standard]
-    binop_eq [of "op \<le>", OF le_number_of_eq refl numeral_1_eq_1, standard]
+    binop_eq [of "op \<le>", OF le_number_of numeral_0_eq_0 refl, standard]
+    binop_eq [of "op \<le>", OF le_number_of numeral_1_eq_1 refl, standard]
+    binop_eq [of "op \<le>", OF le_number_of refl numeral_0_eq_0, standard]
+    binop_eq [of "op \<le>", OF le_number_of refl numeral_1_eq_1, standard]
 
 lemmas arith_special[simp] = 
        add_special diff_special eq_special less_special le_special
--- a/src/HOL/Library/Efficient_Nat.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Library/Efficient_Nat.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -349,7 +349,7 @@
 
 lemma nat_code' [code]:
   "nat (number_of l) = (if neg (number_of l \<Colon> int) then 0 else number_of l)"
-  by auto
+  unfolding nat_number_of_def number_of_is_id neg_def by simp
 
 lemma of_nat_int [code unfold]:
   "of_nat = int" by (simp add: int_def)
--- a/src/HOL/Library/Enum.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Library/Enum.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -365,4 +365,15 @@
 
 end
 
-end
\ No newline at end of file
+instantiation option :: (enum) enum
+begin
+
+definition
+  "enum = None # map Some enum"
+
+instance by default
+  (auto simp add: enum_all enum_option_def, case_tac x, auto intro: inj_onI simp add: distinct_map enum_distinct)
+
+end
+
+end
--- a/src/HOL/Library/Float.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Library/Float.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -499,7 +499,7 @@
 
 lemma int_eq_number_of_eq:
   "(((number_of v)::int)=(number_of w)) = iszero ((number_of (v + uminus w))::int)"
-  by simp
+  by (rule eq_number_of_eq)
 
 lemma int_iszero_number_of_Pls: "iszero (Numeral0::int)"
   by (simp only: iszero_number_of_Pls)
@@ -514,7 +514,7 @@
   by simp
 
 lemma int_less_number_of_eq_neg: "(((number_of x)::int) < number_of y) = neg ((number_of (x + (uminus y)))::int)"
-  unfolding neg_def number_of_is_id by simp
+  by (rule less_number_of_eq_neg)
 
 lemma int_not_neg_number_of_Pls: "\<not> (neg (Numeral0::int))"
   by simp
--- a/src/HOL/Library/Nat_Infinity.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Library/Nat_Infinity.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -146,8 +146,8 @@
   by (simp_all add: plus_inat_def zero_inat_def split: inat.splits)
 
 lemma plus_inat_number [simp]:
-  "(number_of k \<Colon> inat) + number_of l = (if neg (number_of k \<Colon> int) then number_of l
-    else if neg (number_of l \<Colon> int) then number_of k else number_of (k + l))"
+  "(number_of k \<Colon> inat) + number_of l = (if k < Int.Pls then number_of l
+    else if l < Int.Pls then number_of k else number_of (k + l))"
   unfolding number_of_inat_def plus_inat_simps nat_arith(1) if_distrib [symmetric, of _ Fin] ..
 
 lemma iSuc_number [simp]:
@@ -165,6 +165,67 @@
   unfolding iSuc_plus_1 by (simp_all add: add_ac)
 
 
+subsection {* Multiplication *}
+
+instantiation inat :: comm_semiring_1
+begin
+
+definition
+  times_inat_def [code del]:
+  "m * n = (case m of \<infinity> \<Rightarrow> if n = 0 then 0 else \<infinity> | Fin m \<Rightarrow>
+    (case n of \<infinity> \<Rightarrow> if m = 0 then 0 else \<infinity> | Fin n \<Rightarrow> Fin (m * n)))"
+
+lemma times_inat_simps [simp, code]:
+  "Fin m * Fin n = Fin (m * n)"
+  "\<infinity> * \<infinity> = \<infinity>"
+  "\<infinity> * Fin n = (if n = 0 then 0 else \<infinity>)"
+  "Fin m * \<infinity> = (if m = 0 then 0 else \<infinity>)"
+  unfolding times_inat_def zero_inat_def
+  by (simp_all split: inat.split)
+
+instance proof
+  fix a b c :: inat
+  show "(a * b) * c = a * (b * c)"
+    unfolding times_inat_def zero_inat_def
+    by (simp split: inat.split)
+  show "a * b = b * a"
+    unfolding times_inat_def zero_inat_def
+    by (simp split: inat.split)
+  show "1 * a = a"
+    unfolding times_inat_def zero_inat_def one_inat_def
+    by (simp split: inat.split)
+  show "(a + b) * c = a * c + b * c"
+    unfolding times_inat_def zero_inat_def
+    by (simp split: inat.split add: left_distrib)
+  show "0 * a = 0"
+    unfolding times_inat_def zero_inat_def
+    by (simp split: inat.split)
+  show "a * 0 = 0"
+    unfolding times_inat_def zero_inat_def
+    by (simp split: inat.split)
+  show "(0::inat) \<noteq> 1"
+    unfolding zero_inat_def one_inat_def
+    by simp
+qed
+
+end
+
+lemma mult_iSuc: "iSuc m * n = n + m * n"
+  unfolding iSuc_plus_1 by (simp add: ring_simps)
+
+lemma mult_iSuc_right: "m * iSuc n = m + m * n"
+  unfolding iSuc_plus_1 by (simp add: ring_simps)
+
+lemma of_nat_eq_Fin: "of_nat n = Fin n"
+  apply (induct n)
+  apply (simp add: Fin_0)
+  apply (simp add: plus_1_iSuc iSuc_Fin)
+  done
+
+instance inat :: semiring_char_0
+  by default (simp add: of_nat_eq_Fin)
+
+
 subsection {* Ordering *}
 
 instantiation inat :: ordered_ab_semigroup_add
@@ -201,6 +262,15 @@
 
 end
 
+instance inat :: pordered_comm_semiring
+proof
+  fix a b c :: inat
+  assume "a \<le> b" and "0 \<le> c"
+  thus "c * a \<le> c * b"
+    unfolding times_inat_def less_eq_inat_def zero_inat_def
+    by (simp split: inat.splits)
+qed
+
 lemma inat_ord_number [simp]:
   "(number_of m \<Colon> inat) \<le> number_of n \<longleftrightarrow> (number_of m \<Colon> nat) \<le> number_of n"
   "(number_of m \<Colon> inat) < number_of n \<longleftrightarrow> (number_of m \<Colon> nat) < number_of n"
--- a/src/HOL/Library/Numeral_Type.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Library/Numeral_Type.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -14,23 +14,6 @@
 subsection {* Preliminary lemmas *}
 (* These should be moved elsewhere *)
 
-lemma inj_Inl [simp]: "inj_on Inl A"
-  by (rule inj_onI, simp)
-
-lemma inj_Inr [simp]: "inj_on Inr A"
-  by (rule inj_onI, simp)
-
-lemma inj_Some [simp]: "inj_on Some A"
-  by (rule inj_onI, simp)
-
-lemma card_Plus:
-  "[| finite A; finite B |] ==> card (A <+> B) = card A + card B"
-  unfolding Plus_def
-  apply (subgoal_tac "Inl ` A \<inter> Inr ` B = {}")
-  apply (simp add: card_Un_disjoint card_image)
-  apply fast
-  done
-
 lemma (in type_definition) univ:
   "UNIV = Abs ` A"
 proof
--- a/src/HOL/List.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/List.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -3423,7 +3423,7 @@
    (fn n => fn m => Drule.instantiate' [] [SOME n, SOME m] @{thm nibble_pair_of_char.simps})
       nibbles nibbles;
 in
-  PureThy.note_thmss Thm.lemmaK [((Name.binding "nibble_pair_of_char_simps", []), [(thms, [])])]
+  PureThy.note_thmss Thm.lemmaK [((Binding.name "nibble_pair_of_char_simps", []), [(thms, [])])]
   #-> (fn [(_, thms)] => fold_rev Code.add_eqn thms)
 end
 *}
--- a/src/HOL/NatBin.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/NatBin.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -111,8 +111,8 @@
      "int (number_of v) =  
          (if neg (number_of v :: int) then 0  
           else (number_of v :: int))"
-by (simp del: nat_number_of
-	 add: neg_nat nat_number_of_def not_neg_nat add_assoc)
+  unfolding nat_number_of_def number_of_is_id neg_def
+  by simp
 
 
 subsubsection{*Successor *}
@@ -124,10 +124,9 @@
 
 lemma Suc_nat_number_of_add:
      "Suc (number_of v + n) =  
-        (if neg (number_of v :: int) then 1+n else number_of (Int.succ v) + n)" 
-by (simp del: nat_number_of 
-         add: nat_number_of_def neg_nat
-              Suc_nat_eq_nat_zadd1 number_of_succ) 
+        (if neg (number_of v :: int) then 1+n else number_of (Int.succ v) + n)"
+  unfolding nat_number_of_def number_of_is_id neg_def numeral_simps
+  by (simp add: Suc_nat_eq_nat_zadd1 add_ac)
 
 lemma Suc_nat_number_of [simp]:
      "Suc (number_of v) =  
@@ -142,10 +141,11 @@
 (*"neg" is used in rewrite rules for binary comparisons*)
 lemma add_nat_number_of [simp]:
      "(number_of v :: nat) + number_of v' =  
-         (if neg (number_of v :: int) then number_of v'  
-          else if neg (number_of v' :: int) then number_of v  
+         (if v < Int.Pls then number_of v'  
+          else if v' < Int.Pls then number_of v  
           else number_of (v + v'))"
-by (simp add: neg_nat nat_number_of_def nat_add_distrib [symmetric] del: nat_number_of)
+  unfolding nat_number_of_def number_of_is_id numeral_simps
+  by (simp add: nat_add_distrib)
 
 
 subsubsection{*Subtraction *}
@@ -160,19 +160,20 @@
 
 lemma diff_nat_number_of [simp]: 
      "(number_of v :: nat) - number_of v' =  
-        (if neg (number_of v' :: int) then number_of v  
+        (if v' < Int.Pls then number_of v  
          else let d = number_of (v + uminus v') in     
               if neg d then 0 else nat d)"
-by (simp del: nat_number_of add: diff_nat_eq_if nat_number_of_def) 
-
+  unfolding nat_number_of_def number_of_is_id numeral_simps neg_def
+  by auto
 
 
 subsubsection{*Multiplication *}
 
 lemma mult_nat_number_of [simp]:
      "(number_of v :: nat) * number_of v' =  
-       (if neg (number_of v :: int) then 0 else number_of (v * v'))"
-by (simp add: neg_nat nat_number_of_def nat_mult_distrib [symmetric] del: nat_number_of)
+       (if v < Int.Pls then 0 else number_of (v * v'))"
+  unfolding nat_number_of_def number_of_is_id numeral_simps
+  by (simp add: nat_mult_distrib)
 
 
 subsubsection{*Quotient *}
@@ -181,7 +182,8 @@
      "(number_of v :: nat)  div  number_of v' =  
           (if neg (number_of v :: int) then 0  
            else nat (number_of v div number_of v'))"
-by (simp add: neg_nat nat_number_of_def nat_div_distrib [symmetric] del: nat_number_of)
+  unfolding nat_number_of_def number_of_is_id neg_def
+  by (simp add: nat_div_distrib)
 
 lemma one_div_nat_number_of [simp]:
      "Suc 0 div number_of v' = nat (1 div number_of v')" 
@@ -195,7 +197,8 @@
         (if neg (number_of v :: int) then 0  
          else if neg (number_of v' :: int) then number_of v  
          else nat (number_of v mod number_of v'))"
-by (simp add: neg_nat nat_number_of_def nat_mod_distrib [symmetric] del: nat_number_of)
+  unfolding nat_number_of_def number_of_is_id neg_def
+  by (simp add: nat_mod_distrib)
 
 lemma one_mod_nat_number_of [simp]:
      "Suc 0 mod number_of v' =  
@@ -246,28 +249,30 @@
 (*"neg" is used in rewrite rules for binary comparisons*)
 lemma eq_nat_number_of [simp]:
      "((number_of v :: nat) = number_of v') =  
-      (if neg (number_of v :: int) then (iszero (number_of v' :: int) | neg (number_of v' :: int))  
-       else if neg (number_of v' :: int) then iszero (number_of v :: int)  
-       else iszero (number_of (v + uminus v') :: int))"
-apply (simp only: simp_thms neg_nat not_neg_eq_ge_0 nat_number_of_def
-                  eq_nat_nat_iff eq_number_of_eq nat_0 iszero_def
-            split add: split_if cong add: imp_cong)
-apply (simp only: nat_eq_iff nat_eq_iff2)
-apply (simp add: not_neg_eq_ge_0 [symmetric])
-done
+      (if neg (number_of v :: int) then (number_of v' :: int) \<le> 0
+       else if neg (number_of v' :: int) then (number_of v :: int) = 0
+       else v = v')"
+  unfolding nat_number_of_def number_of_is_id neg_def
+  by auto
 
 
 subsubsection{*Less-than (<) *}
 
-(*"neg" is used in rewrite rules for binary comparisons*)
 lemma less_nat_number_of [simp]:
-     "((number_of v :: nat) < number_of v') =  
-         (if neg (number_of v :: int) then neg (number_of (uminus v') :: int)  
-          else neg (number_of (v + uminus v') :: int))"
-  unfolding neg_def nat_number_of_def number_of_is_id
+  "(number_of v :: nat) < number_of v' \<longleftrightarrow>
+    (if v < v' then Int.Pls < v' else False)"
+  unfolding nat_number_of_def number_of_is_id numeral_simps
   by auto
 
 
+subsubsection{*Less-than-or-equal *}
+
+lemma le_nat_number_of [simp]:
+  "(number_of v :: nat) \<le> number_of v' \<longleftrightarrow>
+    (if v \<le> v' then True else v \<le> Int.Pls)"
+  unfolding nat_number_of_def number_of_is_id numeral_simps
+  by auto
+
 (*Maps #n to n for n = 0, 1, 2*)
 lemmas numerals = nat_numeral_0_eq_0 nat_numeral_1_eq_1 numeral_2_eq_2
 
@@ -441,22 +446,21 @@
 text{*Simplification already does @{term "n<0"}, @{term "n\<le>0"} and @{term "0\<le>n"}.*}
 
 lemma eq_number_of_0 [simp]:
-     "(number_of v = (0::nat)) =  
-      (if neg (number_of v :: int) then True else iszero (number_of v :: int))"
-by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric] iszero_0)
+  "number_of v = (0::nat) \<longleftrightarrow> v \<le> Int.Pls"
+  unfolding nat_number_of_def number_of_is_id numeral_simps
+  by auto
 
 lemma eq_0_number_of [simp]:
-     "((0::nat) = number_of v) =  
-      (if neg (number_of v :: int) then True else iszero (number_of v :: int))"
+  "(0::nat) = number_of v \<longleftrightarrow> v \<le> Int.Pls"
 by (rule trans [OF eq_sym_conv eq_number_of_0])
 
 lemma less_0_number_of [simp]:
-     "((0::nat) < number_of v) = neg (number_of (uminus v) :: int)"
-by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric] Pls_def)
-
+   "(0::nat) < number_of v \<longleftrightarrow> Int.Pls < v"
+  unfolding nat_number_of_def number_of_is_id numeral_simps
+  by simp
 
 lemma neg_imp_number_of_eq_0: "neg (number_of v :: int) ==> number_of v = (0::nat)"
-by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric] iszero_0)
+by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric])
 
 
 
@@ -643,26 +647,15 @@
 
 lemma nat_number_of_Bit0:
     "number_of (Int.Bit0 w) = (let n::nat = number_of w in n + n)"
-  apply (simp only: nat_number_of_def Let_def)
-  apply (cases "neg (number_of w :: int)")
-   apply (simp add: neg_nat neg_number_of_Bit0)
-  apply (rule int_int_eq [THEN iffD1])
-  apply (simp only: not_neg_nat neg_number_of_Bit0 int_Suc zadd_int [symmetric] simp_thms)
-  apply (simp only: number_of_Bit0 zadd_assoc)
-  apply simp
-  done
+  unfolding nat_number_of_def number_of_is_id numeral_simps Let_def
+  by auto
 
 lemma nat_number_of_Bit1:
   "number_of (Int.Bit1 w) =
     (if neg (number_of w :: int) then 0
      else let n = number_of w in Suc (n + n))"
-  apply (simp only: nat_number_of_def Let_def split: split_if)
-  apply (intro conjI impI)
-   apply (simp add: neg_nat neg_number_of_Bit1)
-  apply (rule int_int_eq [THEN iffD1])
-  apply (simp only: not_neg_nat neg_number_of_Bit1 int_Suc zadd_int [symmetric] simp_thms)
-  apply (simp only: number_of_Bit1 zadd_assoc)
-  done
+  unfolding nat_number_of_def number_of_is_id numeral_simps neg_def Let_def
+  by auto
 
 lemmas nat_number =
   nat_number_of_Pls nat_number_of_Min
@@ -708,11 +701,12 @@
          (if neg (number_of v :: int) then number_of v' + k  
           else if neg (number_of v' :: int) then number_of v + k  
           else number_of (v + v') + k)"
-by simp
+  unfolding nat_number_of_def number_of_is_id neg_def
+  by auto
 
 lemma nat_number_of_mult_left:
      "number_of v * (number_of v' * (k::nat)) =  
-         (if neg (number_of v :: int) then 0
+         (if v < Int.Pls then 0
           else number_of (v * v') * k)"
 by simp
 
--- a/src/HOL/Nominal/Nominal.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Nominal/Nominal.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,3 @@
-(* $Id$ *)
-
 theory Nominal 
 imports Main Infinite_Set
 uses
--- a/src/HOL/Nominal/nominal_atoms.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Nominal/nominal_atoms.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -170,7 +170,7 @@
     val (swap_eqs, thy3) = fold_map (fn (ak_name, T) => fn thy =>
       let
         val swapT = HOLogic.mk_prodT (T, T) --> T --> T;
-        val swap_name = Sign.full_name thy ("swap_" ^ ak_name);
+        val swap_name = Sign.full_bname thy ("swap_" ^ ak_name);
         val a = Free ("a", T);
         val b = Free ("b", T);
         val c = Free ("c", T);
@@ -199,10 +199,10 @@
     val (prm_eqs, thy4) = fold_map (fn (ak_name, T) => fn thy =>
       let
         val swapT = HOLogic.mk_prodT (T, T) --> T --> T;
-        val swap_name = Sign.full_name thy ("swap_" ^ ak_name)
+        val swap_name = Sign.full_bname thy ("swap_" ^ ak_name)
         val prmT = mk_permT T --> T --> T;
         val prm_name = ak_name ^ "_prm_" ^ ak_name;
-        val qu_prm_name = Sign.full_name thy prm_name;
+        val qu_prm_name = Sign.full_bname thy prm_name;
         val x  = Free ("x", HOLogic.mk_prodT (T, T));
         val xs = Free ("xs", mk_permT T);
         val a  = Free ("a", T) ;
@@ -235,7 +235,7 @@
           val pi = Free ("pi", mk_permT T);
           val a  = Free ("a", T');
           val cperm = Const ("Nominal.perm", mk_permT T --> T' --> T');
-          val cperm_def = Const (Sign.full_name thy' perm_def_name, mk_permT T --> T' --> T');
+          val cperm_def = Const (Sign.full_bname thy' perm_def_name, mk_permT T --> T' --> T');
 
           val name = ak_name ^ "_prm_" ^ ak_name' ^ "_def";
           val def = Logic.mk_equals
@@ -250,7 +250,7 @@
     val (prm_cons_thms,thy6) = 
       thy5 |> PureThy.add_thms (map (fn (ak_name, T) =>
       let
-        val ak_name_qu = Sign.full_name thy5 (ak_name);
+        val ak_name_qu = Sign.full_bname thy5 (ak_name);
         val i_type = Type(ak_name_qu,[]);
         val cat = Const ("Nominal.at",(Term.itselfT i_type)  --> HOLogic.boolT);
         val at_type = Logic.mk_type i_type;
@@ -299,9 +299,9 @@
                         HOLogic.mk_Trueprop (HOLogic.mk_eq (cperm $ pi1 $ x, cperm $ pi2 $ x)));
       in
           AxClass.define_class (cl_name, ["HOL.type"]) []
-                [((Name.binding (cl_name ^ "1"), [Simplifier.simp_add]), [axiom1]),
-                 ((Name.binding (cl_name ^ "2"), []), [axiom2]),                           
-                 ((Name.binding (cl_name ^ "3"), []), [axiom3])] thy
+                [((Binding.name (cl_name ^ "1"), [Simplifier.simp_add]), [axiom1]),
+                 ((Binding.name (cl_name ^ "2"), []), [axiom2]),                           
+                 ((Binding.name (cl_name ^ "3"), []), [axiom3])] thy
       end) ak_names_types thy6;
 
     (* proves that every pt_<ak>-type together with <ak>-type *)
@@ -311,8 +311,8 @@
     val (prm_inst_thms,thy8) = 
       thy7 |> PureThy.add_thms (map (fn (ak_name, T) =>
       let
-        val ak_name_qu = Sign.full_name thy7 ak_name;
-        val pt_name_qu = Sign.full_name thy7 ("pt_"^ak_name);
+        val ak_name_qu = Sign.full_bname thy7 ak_name;
+        val pt_name_qu = Sign.full_bname thy7 ("pt_"^ak_name);
         val i_type1 = TFree("'x",[pt_name_qu]);
         val i_type2 = Type(ak_name_qu,[]);
         val cpt = Const ("Nominal.pt",(Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT);
@@ -338,7 +338,7 @@
      val (fs_ax_classes,thy11) =  fold_map (fn (ak_name, T) => fn thy =>
        let 
           val cl_name = "fs_"^ak_name;
-          val pt_name = Sign.full_name thy ("pt_"^ak_name);
+          val pt_name = Sign.full_bname thy ("pt_"^ak_name);
           val ty = TFree("'a",["HOL.type"]);
           val x   = Free ("x", ty);
           val csupp    = Const ("Nominal.supp", ty --> HOLogic.mk_setT T);
@@ -347,7 +347,7 @@
           val axiom1   = HOLogic.mk_Trueprop (cfinite $ (csupp $ x));
 
        in  
-        AxClass.define_class (cl_name, [pt_name]) [] [((Name.binding (cl_name ^ "1"), []), [axiom1])] thy            
+        AxClass.define_class (cl_name, [pt_name]) [] [((Binding.name (cl_name ^ "1"), []), [axiom1])] thy            
        end) ak_names_types thy8; 
          
      (* proves that every fs_<ak>-type together with <ak>-type   *)
@@ -357,8 +357,8 @@
      val (fs_inst_thms,thy12) = 
        thy11 |> PureThy.add_thms (map (fn (ak_name, T) =>
        let
-         val ak_name_qu = Sign.full_name thy11 ak_name;
-         val fs_name_qu = Sign.full_name thy11 ("fs_"^ak_name);
+         val ak_name_qu = Sign.full_bname thy11 ak_name;
+         val fs_name_qu = Sign.full_bname thy11 ("fs_"^ak_name);
          val i_type1 = TFree("'x",[fs_name_qu]);
          val i_type2 = Type(ak_name_qu,[]);
          val cfs = Const ("Nominal.fs", 
@@ -397,7 +397,7 @@
                                            cperm2 $ (cperm3 $ pi1 $ pi2) $ (cperm1 $ pi1 $ x)));
                in  
                  AxClass.define_class (cl_name, ["HOL.type"]) []
-                   [((Name.binding (cl_name ^ "1"), []), [ax1])] thy'  
+                   [((Binding.name (cl_name ^ "1"), []), [ax1])] thy'  
                end) ak_names_types thy) ak_names_types thy12;
 
         (* proves for every <ak>-combination a cp_<ak1>_<ak2>_inst theorem;     *)
@@ -406,9 +406,9 @@
         val (cp_thms,thy12c) = fold_map (fn (ak_name, T) => fn thy =>
          fold_map (fn (ak_name', T') => fn thy' =>
            let
-             val ak_name_qu  = Sign.full_name thy' (ak_name);
-             val ak_name_qu' = Sign.full_name thy' (ak_name');
-             val cp_name_qu  = Sign.full_name thy' ("cp_"^ak_name^"_"^ak_name');
+             val ak_name_qu  = Sign.full_bname thy' (ak_name);
+             val ak_name_qu' = Sign.full_bname thy' (ak_name');
+             val cp_name_qu  = Sign.full_bname thy' ("cp_"^ak_name^"_"^ak_name');
              val i_type0 = TFree("'a",[cp_name_qu]);
              val i_type1 = Type(ak_name_qu,[]);
              val i_type2 = Type(ak_name_qu',[]);
@@ -442,8 +442,8 @@
           (if not (ak_name = ak_name') 
            then 
                let
-                 val ak_name_qu  = Sign.full_name thy' ak_name;
-                 val ak_name_qu' = Sign.full_name thy' ak_name';
+                 val ak_name_qu  = Sign.full_bname thy' ak_name;
+                 val ak_name_qu' = Sign.full_bname thy' ak_name';
                  val i_type1 = Type(ak_name_qu,[]);
                  val i_type2 = Type(ak_name_qu',[]);
                  val cdj = Const ("Nominal.disjoint",
@@ -487,8 +487,8 @@
      val thy13 = fold (fn ak_name => fn thy =>
         fold (fn ak_name' => fn thy' =>
          let
-           val qu_name =  Sign.full_name thy' ak_name';
-           val cls_name = Sign.full_name thy' ("pt_"^ak_name);
+           val qu_name =  Sign.full_bname thy' ak_name';
+           val cls_name = Sign.full_bname thy' ("pt_"^ak_name);
            val at_inst  = PureThy.get_thm thy' ("at_" ^ ak_name' ^ "_inst");
 
            val proof1 = EVERY [Class.intro_classes_tac [],
@@ -518,7 +518,7 @@
      (* are instances of pt_<ak>        *)
      val thy18 = fold (fn ak_name => fn thy =>
        let
-          val cls_name = Sign.full_name thy ("pt_"^ak_name);
+          val cls_name = Sign.full_bname thy ("pt_"^ak_name);
           val at_thm   = PureThy.get_thm thy ("at_"^ak_name^"_inst");
           val pt_inst  = PureThy.get_thm thy ("pt_"^ak_name^"_inst");
 
@@ -562,8 +562,8 @@
        val thy20 = fold (fn ak_name => fn thy =>
         fold (fn ak_name' => fn thy' =>
         let
-           val qu_name =  Sign.full_name thy' ak_name';
-           val qu_class = Sign.full_name thy' ("fs_"^ak_name);
+           val qu_name =  Sign.full_bname thy' ak_name';
+           val qu_class = Sign.full_bname thy' ("fs_"^ak_name);
            val proof =
                (if ak_name = ak_name'
                 then
@@ -588,7 +588,7 @@
 
        val thy24 = fold (fn ak_name => fn thy => 
         let
-          val cls_name = Sign.full_name thy ("fs_"^ak_name);
+          val cls_name = Sign.full_bname thy ("fs_"^ak_name);
           val fs_inst  = PureThy.get_thm thy ("fs_"^ak_name^"_inst");
           fun fs_proof thm = EVERY [Class.intro_classes_tac [], rtac (thm RS fs1) 1];
 
@@ -628,8 +628,8 @@
          fold (fn ak_name' => fn thy' =>
           fold (fn ak_name'' => fn thy'' =>
             let
-              val name =  Sign.full_name thy'' ak_name;
-              val cls_name = Sign.full_name thy'' ("cp_"^ak_name'^"_"^ak_name'');
+              val name =  Sign.full_bname thy'' ak_name;
+              val cls_name = Sign.full_bname thy'' ("cp_"^ak_name'^"_"^ak_name'');
               val proof =
                 (if (ak_name'=ak_name'') then 
                   (let
@@ -666,7 +666,7 @@
        val thy26 = fold (fn ak_name => fn thy =>
         fold (fn ak_name' => fn thy' =>
         let
-            val cls_name = Sign.full_name thy' ("cp_"^ak_name^"_"^ak_name');
+            val cls_name = Sign.full_bname thy' ("cp_"^ak_name^"_"^ak_name');
             val cp_inst  = PureThy.get_thm thy' ("cp_"^ak_name^"_"^ak_name'^"_inst");
             val pt_inst  = PureThy.get_thm thy' ("pt_"^ak_name^"_inst");
             val at_inst  = PureThy.get_thm thy' ("at_"^ak_name^"_inst");
@@ -698,7 +698,7 @@
           fun discrete_pt_inst discrete_ty defn =
              fold (fn ak_name => fn thy =>
              let
-               val qu_class = Sign.full_name thy ("pt_"^ak_name);
+               val qu_class = Sign.full_bname thy ("pt_"^ak_name);
                val simp_s = HOL_basic_ss addsimps [defn];
                val proof = EVERY [Class.intro_classes_tac [], REPEAT (asm_simp_tac simp_s 1)];
              in 
@@ -708,7 +708,7 @@
           fun discrete_fs_inst discrete_ty defn = 
              fold (fn ak_name => fn thy =>
              let
-               val qu_class = Sign.full_name thy ("fs_"^ak_name);
+               val qu_class = Sign.full_bname thy ("fs_"^ak_name);
                val supp_def = @{thm "Nominal.supp_def"};
                val simp_s = HOL_ss addsimps [supp_def,Collect_const,finite_emptyI,defn];
                val proof = EVERY [Class.intro_classes_tac [], asm_simp_tac simp_s 1];
@@ -719,7 +719,7 @@
           fun discrete_cp_inst discrete_ty defn = 
              fold (fn ak_name' => (fold (fn ak_name => fn thy =>
              let
-               val qu_class = Sign.full_name thy ("cp_"^ak_name^"_"^ak_name');
+               val qu_class = Sign.full_bname thy ("cp_"^ak_name^"_"^ak_name');
                val supp_def = @{thm "Nominal.supp_def"};
                val simp_s = HOL_ss addsimps [defn];
                val proof = EVERY [Class.intro_classes_tac [], asm_simp_tac simp_s 1];
--- a/src/HOL/Nominal/nominal_induct.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Nominal/nominal_induct.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -6,7 +6,7 @@
 
 structure NominalInduct:
 sig
-  val nominal_induct_tac: Proof.context -> (Name.binding option * term) option list list ->
+  val nominal_induct_tac: Proof.context -> (Binding.T option * term) option list list ->
     (string * typ) list -> (string * typ) list list -> thm list ->
     thm list -> int -> RuleCases.cases_tactic
   val nominal_induct_method: Method.src -> Proof.context -> Method.method
--- a/src/HOL/Nominal/nominal_inductive2.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Nominal/nominal_inductive2.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -182,7 +182,7 @@
     fun mk_avoids params name sets =
       let
         val (_, ctxt') = ProofContext.add_fixes_i
-          (map (fn (s, T) => (Name.binding s, SOME T, NoSyn)) params) ctxt;
+          (map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params) ctxt;
         fun mk s =
           let
             val t = Syntax.read_term ctxt' s;
--- a/src/HOL/Nominal/nominal_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Nominal/nominal_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -227,7 +227,7 @@
       map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts';
 
     val ps = map (fn (_, n, _, _) =>
-      (Sign.full_name tmp_thy n, Sign.full_name tmp_thy (n ^ "_Rep"))) dts;
+      (Sign.full_bname tmp_thy n, Sign.full_bname tmp_thy (n ^ "_Rep"))) dts;
     val rps = map Library.swap ps;
 
     fun replace_types (Type ("Nominal.ABS", [T, U])) =
@@ -241,7 +241,7 @@
         map replace_types cargs, NoSyn)) constrs)) dts';
 
     val new_type_names' = map (fn n => n ^ "_Rep") new_type_names;
-    val full_new_type_names' = map (Sign.full_name thy) new_type_names';
+    val full_new_type_names' = map (Sign.full_bname thy) new_type_names';
 
     val ({induction, ...},thy1) =
       DatatypePackage.add_datatype err flat_names new_type_names' dts'' thy;
@@ -263,7 +263,7 @@
     val perm_names' = DatatypeProp.indexify_names (map (fn (i, _) =>
       "perm_" ^ name_of_typ (nth_dtyp i)) descr);
     val perm_names = replicate (length new_type_names) "Nominal.perm" @
-      map (Sign.full_name thy1) (List.drop (perm_names', length new_type_names));
+      map (Sign.full_bname thy1) (List.drop (perm_names', length new_type_names));
     val perm_names_types = perm_names ~~ perm_types;
     val perm_names_types' = perm_names' ~~ perm_types;
 
@@ -289,7 +289,7 @@
               else Const ("Nominal.perm", permT --> T --> T) $ pi $ x
             end;
         in
-          (Attrib.no_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq
+          (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq
             (Free (nth perm_names_types' i) $
                Free ("pi", mk_permT (TFree ("'x", HOLogic.typeS))) $
                list_comb (c, args),
@@ -301,7 +301,7 @@
       PrimrecPackage.add_primrec_overloaded
         (map (fn (s, sT) => (s, sT, false))
            (List.take (perm_names' ~~ perm_names_types, length new_type_names)))
-        (map (fn s => (Name.binding s, NONE, NoSyn)) perm_names') perm_eqs thy1;
+        (map (fn s => (Binding.name s, NONE, NoSyn)) perm_names') perm_eqs thy1;
 
     (**** prove that permutation functions introduced by unfolding are ****)
     (**** equivalent to already existing permutation functions         ****)
@@ -566,16 +566,16 @@
                 (rep_set_name, T))
               end)
                 (descr ~~ rep_set_names))));
-    val rep_set_names'' = map (Sign.full_name thy3) rep_set_names';
+    val rep_set_names'' = map (Sign.full_bname thy3) rep_set_names';
 
     val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy4) =
         InductivePackage.add_inductive_global (serial_string ())
           {quiet_mode = false, verbose = false, kind = Thm.internalK,
-           alt_name = Name.binding big_rep_name, coind = false, no_elim = true, no_ind = false,
+           alt_name = Binding.name big_rep_name, coind = false, no_elim = true, no_ind = false,
            skip_mono = true}
-          (map (fn (s, T) => ((Name.binding s, T --> HOLogic.boolT), NoSyn))
+          (map (fn (s, T) => ((Binding.name s, T --> HOLogic.boolT), NoSyn))
              (rep_set_names' ~~ recTs'))
-          [] (map (fn x => (Attrib.no_binding, x)) intr_ts) [] thy3;
+          [] (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] thy3;
 
     (**** Prove that representing set is closed under permutation ****)
 
@@ -768,7 +768,7 @@
     val newTs' = Library.take (length new_type_names, recTs');
     val newTs = Library.take (length new_type_names, recTs);
 
-    val full_new_type_names = map (Sign.full_name thy) new_type_names;
+    val full_new_type_names = map (Sign.full_bname thy) new_type_names;
 
     fun make_constr_def tname T T' ((thy, defs, eqns),
         (((cname_rep, _), (cname, cargs)), (cname', mx))) =
@@ -1432,7 +1432,7 @@
       if length descr'' = 1 then [big_rec_name] else
         map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
           (1 upto (length descr''));
-    val rec_set_names =  map (Sign.full_name thy10) rec_set_names';
+    val rec_set_names =  map (Sign.full_bname thy10) rec_set_names';
 
     val rec_fns = map (uncurry (mk_Free "f"))
       (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts)));
@@ -1509,12 +1509,12 @@
       thy10 |>
         InductivePackage.add_inductive_global (serial_string ())
           {quiet_mode = ! quiet_mode, verbose = false, kind = Thm.internalK,
-           alt_name = Name.binding big_rec_name, coind = false, no_elim = false, no_ind = false,
+           alt_name = Binding.name big_rec_name, coind = false, no_elim = false, no_ind = false,
            skip_mono = true}
-          (map (fn (s, T) => ((Name.binding s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
+          (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
           (map dest_Free rec_fns)
-          (map (fn x => (Attrib.no_binding, x)) rec_intr_ts) [] ||>
-      PureThy.hide_fact true (NameSpace.append (Sign.full_name thy10 big_rec_name) "induct");
+          (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] ||>
+      PureThy.hide_fact true (NameSpace.append (Sign.full_bname thy10 big_rec_name) "induct");
 
     (** equivariance **)
 
@@ -2002,7 +2002,7 @@
     (* define primrec combinators *)
 
     val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
-    val reccomb_names = map (Sign.full_name thy11)
+    val reccomb_names = map (Sign.full_bname thy11)
       (if length descr'' = 1 then [big_reccomb_name] else
         (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
           (1 upto (length descr''))));
--- a/src/HOL/Nominal/nominal_primrec.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Nominal/nominal_primrec.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -9,13 +9,13 @@
 signature NOMINAL_PRIMREC =
 sig
   val add_primrec: string -> string list option -> string option ->
-    ((Name.binding * string) * Attrib.src list) list -> theory -> Proof.state
+    ((Binding.T * string) * Attrib.src list) list -> theory -> Proof.state
   val add_primrec_unchecked: string -> string list option -> string option ->
-    ((Name.binding * string) * Attrib.src list) list -> theory -> Proof.state
+    ((Binding.T * string) * Attrib.src list) list -> theory -> Proof.state
   val add_primrec_i: string -> term list option -> term option ->
-    ((Name.binding * term) * attribute list) list -> theory -> Proof.state
+    ((Binding.T * term) * attribute list) list -> theory -> Proof.state
   val add_primrec_unchecked_i: string -> term list option -> term option ->
-    ((Name.binding * term) * attribute list) list -> theory -> Proof.state
+    ((Binding.T * term) * attribute list) list -> theory -> Proof.state
 end;
 
 structure NominalPrimrec : NOMINAL_PRIMREC =
@@ -230,7 +230,7 @@
 fun gen_primrec_i note def alt_name invs fctxt eqns_atts thy =
   let
     val (raw_eqns, atts) = split_list eqns_atts;
-    val eqns = map (apfst Name.name_of) raw_eqns;
+    val eqns = map (apfst Binding.base_name) raw_eqns;
     val dt_info = NominalPackage.get_nominal_datatypes thy;
     val rec_eqns = fold_rev (process_eqn thy o snd) eqns [];
     val lsrs :: lsrss = maps (fn (_, (_, _, eqns)) =>
--- a/src/HOL/Presburger.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Presburger.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -411,7 +411,7 @@
   by (simp cong: conj_cong)
 lemma int_eq_number_of_eq:
   "(((number_of v)::int) = (number_of w)) = iszero ((number_of (v + (uminus w)))::int)"
-  by simp
+  by (rule eq_number_of_eq)
 
 lemma mod_eq0_dvd_iff[presburger]: "(m::nat) mod n = 0 \<longleftrightarrow> n dvd m"
 unfolding dvd_eq_mod_eq_0[symmetric] ..
--- a/src/HOL/ROOT.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/ROOT.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -3,7 +3,7 @@
 Classical Higher-order Logic -- batteries included.
 *)
 
-use_thy "Complex/Complex_Main";
+use_thy "Complex_Main";
 
 val HOL_proofs = ! Proofterm.proofs;
 
--- a/src/HOL/Statespace/state_space.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Statespace/state_space.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -268,7 +268,7 @@
                 (map (fn n => Locale.Rename (Locale.Locale (Locale.intern thy "var")
                                             ,[SOME (n,NONE)])) all_comps);
 
-    val full_name = Sign.full_name thy name;
+    val full_name = Sign.full_bname thy name;
     val dist_thm_name = distinct_compsN;
     val dist_thm_full_name =
         let val prefix = fold1' (fn name => fn prfx => prfx ^ "_" ^ name) all_comps "";
@@ -302,7 +302,7 @@
 
     val attr = Attrib.internal type_attr;
 
-    val assumes = Element.Assumes [((Name.binding dist_thm_name,[attr]),
+    val assumes = Element.Assumes [((Binding.name dist_thm_name,[attr]),
                     [(HOLogic.Trueprop $
                       (Const ("DistinctTreeProver.all_distinct",
                                  Type ("DistinctTreeProver.tree",[nameT]) --> HOLogic.boolT) $
@@ -373,7 +373,7 @@
 
 fun statespace_definition state_type args name parents parent_comps components thy =
   let
-    val full_name = Sign.full_name thy name;
+    val full_name = Sign.full_bname thy name;
     val all_comps = parent_comps @ components;
 
     val components' = map (fn (n,T) => (n,(T,full_name))) components;
@@ -443,7 +443,7 @@
          NONE => []
        | SOME s =>
           let
-	    val fx = Element.Fixes [(Name.binding s,SOME (string_of_typ stateT),NoSyn)];
+	    val fx = Element.Fixes [(Binding.name s,SOME (string_of_typ stateT),NoSyn)];
             val cs = Element.Constrains
                        (map (fn (n,T) =>  (n,string_of_typ T))
                          ((map (fn (n,_) => (n,nameT)) all_comps) @
@@ -490,7 +490,7 @@
 
     fun add_parent (Ts,pname,rs) env =
       let
-        val full_pname = Sign.full_name thy pname;
+        val full_pname = Sign.full_bname thy pname;
         val {args,components,...} =
               (case get_statespace (Context.Theory thy) full_pname of
                 SOME r => r
--- a/src/HOL/Sum_Type.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Sum_Type.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -93,16 +93,17 @@
 lemma Inr_Rep_inject: "Inr_Rep(b) = Inr_Rep(d) ==> b=d"
 by (auto simp add: Inr_Rep_def expand_fun_eq)
 
-lemma inj_Inl: "inj(Inl)"
+lemma inj_Inl [simp]: "inj_on Inl A"
 apply (simp add: Inl_def)
 apply (rule inj_onI)
 apply (erule inj_on_Abs_Sum [THEN inj_onD, THEN Inl_Rep_inject])
 apply (rule Inl_RepI)
 apply (rule Inl_RepI)
 done
+
 lemmas Inl_inject = inj_Inl [THEN injD, standard]
 
-lemma inj_Inr: "inj(Inr)"
+lemma inj_Inr [simp]: "inj_on Inr A"
 apply (simp add: Inr_def)
 apply (rule inj_onI)
 apply (erule inj_on_Abs_Sum [THEN inj_onD, THEN Inr_Rep_inject])
--- a/src/HOL/Tools/ComputeNumeral.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/ComputeNumeral.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -116,7 +116,7 @@
 lemma nat_norm_number_of: "nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)"
   apply (simp add: nat_norm_number_of_def)
   unfolding lezero_def iszero_def neg_def
-  apply (simp add: number_of_is_id)
+  apply (simp add: numeral_simps)
   done
 
 (* Normalization of nat literals *)
@@ -129,7 +129,8 @@
 
 (* Addition for nat *)
 lemma natadd: "number_of x + ((number_of y)::nat) = (if neg x then (number_of y) else (if neg y then number_of x else (number_of (x + y))))"
-  by (auto simp add: number_of_is_id)
+  unfolding nat_number_of_def number_of_is_id neg_def
+  by auto
 
 (* Subtraction for nat *)
 lemma natsub: "(number_of x) - ((number_of y)::nat) = 
@@ -140,17 +141,14 @@
 (* Multiplication for nat *)
 lemma natmul: "(number_of x) * ((number_of y)::nat) = 
   (if neg x then 0 else (if neg y then 0 else number_of (x * y)))"
-  apply (auto simp add: number_of_is_id neg_def iszero_def)
-  apply (case_tac "x > 0")
-  apply auto
-  apply (simp add: mult_strict_left_mono[where a=y and b=0 and c=x, simplified])
-  done
+  unfolding nat_number_of_def number_of_is_id neg_def
+  by (simp add: nat_mult_distrib)
 
 lemma nateq: "(((number_of x)::nat) = (number_of y)) = ((lezero x \<and> lezero y) \<or> (x = y))"
   by (auto simp add: iszero_def lezero_def neg_def number_of_is_id)
 
 lemma natless: "(((number_of x)::nat) < (number_of y)) = ((x < y) \<and> (\<not> (lezero y)))"
-  by (auto simp add: number_of_is_id neg_def lezero_def)
+  by (simp add: lezero_def numeral_simps not_le)
 
 lemma natle: "(((number_of x)::nat) \<le> (number_of y)) = (y < x \<longrightarrow> lezero x)"
   by (auto simp add: number_of_is_id lezero_def nat_number_of_def)
--- a/src/HOL/Tools/datatype_abs_proofs.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/datatype_abs_proofs.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -108,7 +108,7 @@
       if length descr' = 1 then [big_rec_name'] else
         map ((curry (op ^) (big_rec_name' ^ "_")) o string_of_int)
           (1 upto (length descr'));
-    val rec_set_names = map (Sign.full_name thy0) rec_set_names';
+    val rec_set_names = map (Sign.full_bname thy0) rec_set_names';
 
     val (rec_result_Ts, reccomb_fn_Ts) = DatatypeProp.make_primrec_Ts descr sorts used;
 
@@ -156,11 +156,11 @@
     val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
         InductivePackage.add_inductive_global (serial_string ())
           {quiet_mode = ! quiet_mode, verbose = false, kind = Thm.internalK,
-            alt_name = Name.binding big_rec_name', coind = false, no_elim = false, no_ind = true,
+            alt_name = Binding.name big_rec_name', coind = false, no_elim = false, no_ind = true,
             skip_mono = true}
-          (map (fn (s, T) => ((Name.binding s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
+          (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
           (map dest_Free rec_fns)
-          (map (fn x => (Attrib.no_binding, x)) rec_intr_ts) [] thy0;
+          (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] thy0;
 
     (* prove uniqueness and termination of primrec combinators *)
 
@@ -226,7 +226,7 @@
     (* define primrec combinators *)
 
     val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
-    val reccomb_names = map (Sign.full_name thy1)
+    val reccomb_names = map (Sign.full_bname thy1)
       (if length descr' = 1 then [big_reccomb_name] else
         (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
           (1 upto (length descr'))));
@@ -294,7 +294,7 @@
       in Const (@{const_name undefined}, Ts @ Ts' ---> T')
       end) constrs) descr';
 
-    val case_names = map (fn s => Sign.full_name thy1 (s ^ "_case")) new_type_names;
+    val case_names = map (fn s => Sign.full_bname thy1 (s ^ "_case")) new_type_names;
 
     (* define case combinators via primrec combinators *)
 
@@ -317,7 +317,7 @@
           val fns = (List.concat (Library.take (i, case_dummy_fns))) @
             fns2 @ (List.concat (Library.drop (i + 1, case_dummy_fns)));
           val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T');
-          val decl = ((Name.binding (Sign.base_name name), caseT), NoSyn);
+          val decl = ((Binding.name (Sign.base_name name), caseT), NoSyn);
           val def = ((Sign.base_name name) ^ "_def",
             Logic.mk_equals (list_comb (Const (name, caseT), fns1),
               list_comb (reccomb, (List.concat (Library.take (i, case_dummy_fns))) @
--- a/src/HOL/Tools/datatype_codegen.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/datatype_codegen.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -440,7 +440,7 @@
           (mk_side @{const_name eq_class.eq}, mk_side @{const_name "op ="}));
         val def' = Syntax.check_term lthy def;
         val ((_, (_, thm)), lthy') = Specification.definition
-          (NONE, (Attrib.no_binding, def')) lthy;
+          (NONE, (Attrib.empty_binding, def')) lthy;
         val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy);
         val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm;
       in (thm', lthy') end;
--- a/src/HOL/Tools/datatype_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/datatype_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -567,7 +567,7 @@
 
     val (tyvars, _, _, _)::_ = dts;
     val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
-      let val full_tname = Sign.full_name tmp_thy (Syntax.type_name tname mx)
+      let val full_tname = Sign.full_bname tmp_thy (Syntax.type_name tname mx)
       in (case duplicates (op =) tvs of
             [] => if eq_set (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
                   else error ("Mutually recursive datatypes must have same type parameters")
@@ -586,8 +586,8 @@
             val _ = (case fold (curry add_typ_tfree_names) cargs' [] \\ tvs of
                 [] => ()
               | vs => error ("Extra type variables on rhs: " ^ commas vs))
-          in (constrs @ [((if flat_names then Sign.full_name tmp_thy else
-                Sign.full_name_path tmp_thy tname) (Syntax.const_name cname mx'),
+          in (constrs @ [((if flat_names then Sign.full_bname tmp_thy else
+                Sign.full_bname_path tmp_thy tname) (Syntax.const_name cname mx'),
                    map (dtyp_of_typ new_dts) cargs')],
               constr_syntax' @ [(cname, mx')], sorts'')
           end handle ERROR msg =>
@@ -600,7 +600,7 @@
       in
         case duplicates (op =) (map fst constrs') of
            [] =>
-             (dts' @ [(i, (Sign.full_name tmp_thy (Syntax.type_name tname mx),
+             (dts' @ [(i, (Sign.full_bname tmp_thy (Syntax.type_name tname mx),
                 map DtTFree tvs, constrs'))],
               constr_syntax @ [constr_syntax'], sorts', i + 1)
          | dups => error ("Duplicate constructors " ^ commas dups ^
--- a/src/HOL/Tools/datatype_rep_proofs.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/datatype_rep_proofs.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -77,7 +77,7 @@
       (if length descr' = 1 then [big_rec_name] else
         (map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
           (1 upto (length descr'))));
-    val rep_set_names = map (Sign.full_name thy1) rep_set_names';
+    val rep_set_names = map (Sign.full_bname thy1) rep_set_names';
 
     val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr);
     val leafTs' = get_nonrec_types descr' sorts;
@@ -184,10 +184,10 @@
     val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) =
         InductivePackage.add_inductive_global (serial_string ())
           {quiet_mode = ! quiet_mode, verbose = false, kind = Thm.internalK,
-           alt_name = Name.binding big_rec_name, coind = false, no_elim = true, no_ind = false,
+           alt_name = Binding.name big_rec_name, coind = false, no_elim = true, no_ind = false,
            skip_mono = true}
-          (map (fn s => ((Name.binding s, UnivT'), NoSyn)) rep_set_names') []
-          (map (fn x => (Attrib.no_binding, x)) intr_ts) [] thy1;
+          (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') []
+          (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] thy1;
 
     (********************************* typedef ********************************)
 
@@ -210,7 +210,7 @@
     val rep_names' = map (fn i => big_rep_name ^ (string_of_int i))
       (1 upto (length (flat (tl descr))));
     val all_rep_names = map (Sign.intern_const thy3) rep_names @
-      map (Sign.full_name thy3) rep_names';
+      map (Sign.full_bname thy3) rep_names';
 
     (* isomorphism declarations *)
 
--- a/src/HOL/Tools/function_package/fundef_common.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/function_package/fundef_common.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -82,7 +82,7 @@
                                       psimps, pinducts, termination, defname}) phi =
     let
       val term = Morphism.term phi val thm = Morphism.thm phi val fact = Morphism.fact phi
-      val name = Name.name_of o Morphism.name phi o Name.binding
+      val name = Binding.base_name o Morphism.binding phi o Binding.name
     in
       FundefCtxData { add_simps = add_simps, case_names = case_names,
                       fs = map term fs, R = term R, psimps = fact psimps, 
--- a/src/HOL/Tools/function_package/fundef_core.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/function_package/fundef_core.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -485,7 +485,7 @@
               |> Syntax.check_term lthy
 
       val ((f, (_, f_defthm)), lthy) =
-        LocalTheory.define Thm.internalK ((Name.binding (function_name fname), mixfix), ((Name.binding fdefname, []), f_def)) lthy
+        LocalTheory.define Thm.internalK ((Binding.name (function_name fname), mixfix), ((Binding.name fdefname, []), f_def)) lthy
     in
       ((f, f_defthm), lthy)
     end
@@ -898,7 +898,7 @@
           PROFILE "def_rel" (define_recursion_relation (rel_name defname) domT ranT fvar f abstract_qglrs clauses RCss) lthy
 
       val (_, lthy) =
-          LocalTheory.abbrev Syntax.mode_default ((Name.binding (dom_name defname), NoSyn), mk_acc domT R) lthy
+          LocalTheory.abbrev Syntax.mode_default ((Binding.name (dom_name defname), NoSyn), mk_acc domT R) lthy
 
       val newthy = ProofContext.theory_of lthy
       val clauses = map (transfer_clause_ctx newthy) clauses
--- a/src/HOL/Tools/function_package/fundef_lib.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/function_package/fundef_lib.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -56,7 +56,7 @@
 fun dest_all_all_ctx ctx (Const ("all", _) $ Abs (a as (n,T,b))) =
     let
       val [(n', _)] = Variable.variant_frees ctx [] [(n,T)]
-      val (_, ctx') = ProofContext.add_fixes_i [(Name.binding n', SOME T, NoSyn)] ctx
+      val (_, ctx') = ProofContext.add_fixes_i [(Binding.name n', SOME T, NoSyn)] ctx
 
       val (n'', body) = Term.dest_abs (n', T, b) 
       val _ = (n' = n'') orelse error "dest_all_ctx"
--- a/src/HOL/Tools/function_package/fundef_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/function_package/fundef_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -9,14 +9,14 @@
 
 signature FUNDEF_PACKAGE =
 sig
-    val add_fundef :  (Name.binding * string option * mixfix) list
+    val add_fundef :  (Binding.T * string option * mixfix) list
                       -> (Attrib.binding * string) list 
                       -> FundefCommon.fundef_config
                       -> bool list
                       -> local_theory
                       -> Proof.state
 
-    val add_fundef_i:  (Name.binding * typ option * mixfix) list
+    val add_fundef_i:  (Binding.T * typ option * mixfix) list
                        -> (Attrib.binding * term) list
                        -> FundefCommon.fundef_config
                        -> bool list
@@ -36,7 +36,7 @@
 open FundefLib
 open FundefCommon
 
-fun note_theorem ((name, atts), ths) = LocalTheory.note Thm.theoremK ((Name.binding name, atts), ths)
+fun note_theorem ((name, atts), ths) = LocalTheory.note Thm.theoremK ((Binding.name name, atts), ths)
 
 fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_" 
 
@@ -96,8 +96,8 @@
 fun gen_add_fundef is_external prep fixspec eqnss config flags lthy =
     let
       val ((fixes0, spec0), ctxt') = prep fixspec (map (fn (n_a, eq) => [(n_a, [eq])]) eqnss) lthy
-      val fixes = map (apfst (apfst Name.name_of)) fixes0;
-      val spec = map (apfst (apfst Name.name_of)) spec0;
+      val fixes = map (apfst (apfst Binding.base_name)) fixes0;
+      val spec = map (apfst (apfst Binding.base_name)) spec0;
       val (eqs, post, sort_cont, cnames) = FundefCommon.get_preproc lthy config flags ctxt' fixes spec
 
       val defname = mk_defname fixes
@@ -147,10 +147,10 @@
         val goal = HOLogic.mk_Trueprop (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
     in
       lthy
-        |> ProofContext.note_thmss_i "" [((Name.no_binding, [ContextRules.rule_del]), [([allI], [])])] |> snd
-        |> ProofContext.note_thmss_i "" [((Name.no_binding, [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
+        |> ProofContext.note_thmss_i "" [((Binding.empty, [ContextRules.rule_del]), [([allI], [])])] |> snd
+        |> ProofContext.note_thmss_i "" [((Binding.empty, [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
         |> ProofContext.note_thmss_i ""
-          [((Name.binding "termination", [ContextRules.intro_bang (SOME 0)]),
+          [((Binding.name "termination", [ContextRules.intro_bang (SOME 0)]),
             [([Goal.norm_result termination], [])])] |> snd
         |> Proof.theorem_i NONE (total_termination_afterqed data) [[(goal, [])]]
     end
--- a/src/HOL/Tools/function_package/inductive_wrap.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/function_package/inductive_wrap.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -44,14 +44,14 @@
             {quiet_mode = false,
               verbose = ! Toplevel.debug,
               kind = Thm.internalK,
-              alt_name = Name.no_binding,
+              alt_name = Binding.empty,
               coind = false,
               no_elim = false,
               no_ind = false,
               skip_mono = true}
-            [((Name.binding R, T), NoSyn)] (* the relation *)
+            [((Binding.name R, T), NoSyn)] (* the relation *)
             [] (* no parameters *)
-            (map (fn t => (Attrib.no_binding, t)) defs) (* the intros *)
+            (map (fn t => (Attrib.empty_binding, t)) defs) (* the intros *)
             [] (* no special monos *)
             lthy
 
--- a/src/HOL/Tools/function_package/mutual.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/function_package/mutual.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -150,8 +150,8 @@
       fun def ((MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs, f_def, ...}), (_, mixfix)) lthy =
           let
             val ((f, (_, f_defthm)), lthy') =
-              LocalTheory.define Thm.internalK ((Name.binding fname, mixfix),
-                                            ((Name.binding (fname ^ "_def"), []), Term.subst_bound (fsum, f_def)))
+              LocalTheory.define Thm.internalK ((Binding.name fname, mixfix),
+                                            ((Binding.name (fname ^ "_def"), []), Term.subst_bound (fsum, f_def)))
                               lthy
           in
             (MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs=cargTs, f_def=f_def,
--- a/src/HOL/Tools/function_package/size.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/function_package/size.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -89,7 +89,7 @@
       map (fn (T as Type (s, _), optname) =>
         let
           val s' = the_default (Sign.base_name s) optname ^ "_size";
-          val s'' = Sign.full_name thy s'
+          val s'' = Sign.full_bname thy s'
         in
           (s'',
            (list_comb (Const (s'', param_size_fTs @ [T] ---> HOLogic.natT),
@@ -133,7 +133,7 @@
       let
         val (Free (c, _), rhs) = (Logic.dest_equals o Syntax.check_term lthy) eq;
         val ((_, (_, thm)), lthy') = lthy |> LocalTheory.define Thm.definitionK
-          ((Name.binding c, NoSyn), ((Name.binding def_name, []), rhs));
+          ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs));
         val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy');
         val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm;
       in (thm', lthy') end;
--- a/src/HOL/Tools/inductive_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/inductive_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/inductive_package.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Author:     Stefan Berghofer and Markus Wenzel, TU Muenchen
 
@@ -39,17 +38,17 @@
     thm list list * local_theory
   type inductive_flags
   val add_inductive_i:
-    inductive_flags -> ((Name.binding * typ) * mixfix) list ->
+    inductive_flags -> ((Binding.T * typ) * mixfix) list ->
     (string * typ) list -> (Attrib.binding * term) list -> thm list -> local_theory ->
     inductive_result * local_theory
   val add_inductive: bool -> bool ->
-    (Name.binding * string option * mixfix) list ->
-    (Name.binding * string option * mixfix) list ->
+    (Binding.T * string option * mixfix) list ->
+    (Binding.T * string option * mixfix) list ->
     (Attrib.binding * string) list ->
     (Facts.ref * Attrib.src list) list ->
     local_theory -> inductive_result * local_theory
   val add_inductive_global: string -> inductive_flags ->
-    ((Name.binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
+    ((Binding.T * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
     thm list -> theory -> inductive_result * theory
   val arities_of: thm -> (string * int) list
   val params_of: thm -> term list
@@ -64,16 +63,16 @@
 sig
   include BASIC_INDUCTIVE_PACKAGE
   type add_ind_def
-  val declare_rules: string -> Name.binding -> bool -> bool -> string list ->
-    thm list -> Name.binding list -> Attrib.src list list -> (thm * string list) list ->
+  val declare_rules: string -> Binding.T -> bool -> bool -> string list ->
+    thm list -> Binding.T list -> Attrib.src list list -> (thm * string list) list ->
     thm -> local_theory -> thm list * thm list * thm * local_theory
   val add_ind_def: add_ind_def
   val gen_add_inductive_i: add_ind_def -> inductive_flags ->
-    ((Name.binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
+    ((Binding.T * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
     thm list -> local_theory -> inductive_result * local_theory
   val gen_add_inductive: add_ind_def -> bool -> bool ->
-    (Name.binding * string option * mixfix) list ->
-    (Name.binding * string option * mixfix) list ->
+    (Binding.T * string option * mixfix) list ->
+    (Binding.T * string option * mixfix) list ->
     (Attrib.binding * string) list -> (Facts.ref * Attrib.src list) list ->
     local_theory -> inductive_result * local_theory
   val gen_ind_decl: add_ind_def -> bool ->
@@ -261,7 +260,7 @@
 
 fun check_rule ctxt cs params ((binding, att), rule) =
   let
-    val name = Name.name_of binding;
+    val err_name = Binding.display binding;
     val params' = Term.variant_frees rule (Logic.strip_params rule);
     val frees = rev (map Free params');
     val concl = subst_bounds (frees, Logic.strip_assums_concl rule);
@@ -279,7 +278,7 @@
 
     fun check_prem' prem t =
       if head_of t mem cs then
-        check_ind (err_in_prem ctxt name rule prem) t
+        check_ind (err_in_prem ctxt err_name rule prem) t
       else (case t of
           Abs (_, _, t) => check_prem' prem t
         | t $ u => (check_prem' prem t; check_prem' prem u)
@@ -287,15 +286,15 @@
 
     fun check_prem (prem, aprem) =
       if can HOLogic.dest_Trueprop aprem then check_prem' prem prem
-      else err_in_prem ctxt name rule prem "Non-atomic premise";
+      else err_in_prem ctxt err_name rule prem "Non-atomic premise";
   in
     (case concl of
        Const ("Trueprop", _) $ t =>
          if head_of t mem cs then
-           (check_ind (err_in_rule ctxt name rule') t;
+           (check_ind (err_in_rule ctxt err_name rule') t;
             List.app check_prem (prems ~~ aprems))
-         else err_in_rule ctxt name rule' bad_concl
-     | _ => err_in_rule ctxt name rule' bad_concl);
+         else err_in_rule ctxt err_name rule' bad_concl
+     | _ => err_in_rule ctxt err_name rule' bad_concl);
     ((binding, att), arule)
   end;
 
@@ -638,14 +637,14 @@
     (* add definiton of recursive predicates to theory *)
 
     val rec_name =
-      if Binding.is_nothing alt_name then
-        Binding.binding (space_implode "_" (map (Name.name_of o fst) cnames_syn))
+      if Binding.is_empty alt_name then
+        Binding.name (space_implode "_" (map (Binding.base_name o fst) cnames_syn))
       else alt_name;
 
     val ((rec_const, (_, fp_def)), ctxt') = ctxt |>
       LocalTheory.define Thm.internalK
         ((rec_name, case cnames_syn of [(_, syn)] => syn | _ => NoSyn),
-         (Attrib.no_binding, fold_rev lambda params
+         (Attrib.empty_binding, fold_rev lambda params
            (Const (fp_name, (predT --> predT) --> predT) $ fp_fun)));
     val fp_def' = Simplifier.rewrite (HOL_basic_ss addsimps [fp_def])
       (cterm_of (ProofContext.theory_of ctxt') (list_comb (rec_const, params)));
@@ -656,7 +655,7 @@
           val xs = map Free (Variable.variant_frees ctxt intr_ts
             (mk_names "x" (length Ts) ~~ Ts))
         in
-          (name_mx, (Attrib.no_binding, fold_rev lambda (params @ xs)
+          (name_mx, (Attrib.empty_binding, fold_rev lambda (params @ xs)
             (list_comb (rec_const, params @ make_bool_args' bs i @
               make_args argTs (xs ~~ Ts)))))
         end) (cnames_syn ~~ cs);
@@ -672,9 +671,9 @@
 fun declare_rules kind rec_binding coind no_ind cnames intrs intr_bindings intr_atts
       elims raw_induct ctxt =
   let
-    val rec_name = Name.name_of rec_binding;
-    val rec_qualified = Name.qualified rec_name;
-    val intr_names = map Name.name_of intr_bindings;
+    val rec_name = Binding.base_name rec_binding;
+    val rec_qualified = Binding.qualify rec_name;
+    val intr_names = map Binding.base_name intr_bindings;
     val ind_case_names = RuleCases.case_names intr_names;
     val induct =
       if coind then
@@ -693,23 +692,23 @@
       map (hd o snd);
     val (((_, elims'), (_, [induct'])), ctxt2) =
       ctxt1 |>
-      LocalTheory.note kind ((rec_qualified (Name.binding "intros"), []), intrs') ||>>
+      LocalTheory.note kind ((rec_qualified (Binding.name "intros"), []), intrs') ||>>
       fold_map (fn (name, (elim, cases)) =>
-        LocalTheory.note kind ((Name.binding (NameSpace.qualified (Sign.base_name name) "cases"),
+        LocalTheory.note kind ((Binding.name (NameSpace.qualified (Sign.base_name name) "cases"),
           [Attrib.internal (K (RuleCases.case_names cases)),
            Attrib.internal (K (RuleCases.consumes 1)),
            Attrib.internal (K (Induct.cases_pred name)),
            Attrib.internal (K (ContextRules.elim_query NONE))]), [elim]) #>
         apfst (hd o snd)) (if null elims then [] else cnames ~~ elims) ||>>
       LocalTheory.note kind
-        ((rec_qualified (Name.binding (coind_prefix coind ^ "induct")),
+        ((rec_qualified (Binding.name (coind_prefix coind ^ "induct")),
           map (Attrib.internal o K) (#2 induct)), [rulify (#1 induct)]);
 
     val ctxt3 = if no_ind orelse coind then ctxt2 else
       let val inducts = cnames ~~ ProjectRule.projects ctxt2 (1 upto length cnames) induct'
       in
         ctxt2 |>
-        LocalTheory.notes kind [((rec_qualified (Name.binding "inducts"), []),
+        LocalTheory.notes kind [((rec_qualified (Binding.name "inducts"), []),
           inducts |> map (fn (name, th) => ([th],
             [Attrib.internal (K ind_case_names),
              Attrib.internal (K (RuleCases.consumes 1)),
@@ -718,24 +717,24 @@
   in (intrs', elims', induct', ctxt3) end;
 
 type inductive_flags =
-  {quiet_mode: bool, verbose: bool, kind: string, alt_name: Name.binding,
+  {quiet_mode: bool, verbose: bool, kind: string, alt_name: Binding.T,
    coind: bool, no_elim: bool, no_ind: bool, skip_mono: bool}
 
 type add_ind_def =
   inductive_flags ->
   term list -> (Attrib.binding * term) list -> thm list ->
-  term list -> (Name.binding * mixfix) list ->
+  term list -> (Binding.T * mixfix) list ->
   local_theory -> inductive_result * local_theory
 
 fun add_ind_def {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono}
     cs intros monos params cnames_syn ctxt =
   let
     val _ = null cnames_syn andalso error "No inductive predicates given";
-    val names = map (Name.name_of o fst) cnames_syn;
+    val names = map (Binding.base_name o fst) cnames_syn;
     val _ = message (quiet_mode andalso not verbose)
       ("Proofs for " ^ coind_prefix coind ^ "inductive predicate(s) " ^ commas_quote names);
 
-    val cnames = map (Sign.full_name (ProofContext.theory_of ctxt) o Name.name_of o #1) cnames_syn;  (* FIXME *)
+    val cnames = map (Sign.full_name (ProofContext.theory_of ctxt) o #1) cnames_syn;  (* FIXME *)
     val ((intr_names, intr_atts), intr_ts) =
       apfst split_list (split_list (map (check_rule ctxt cs params) intros));
 
@@ -746,7 +745,7 @@
     val (intrs, unfold) = prove_intrs quiet_mode coind mono fp_def (length bs + length xs)
       params intr_ts rec_preds_defs ctxt1;
     val elims = if no_elim then [] else
-      prove_elims quiet_mode cs params intr_ts (map Name.name_of intr_names)
+      prove_elims quiet_mode cs params intr_ts (map Binding.base_name intr_names)
         unfold rec_preds_defs ctxt1;
     val raw_induct = zero_var_indexes
       (if no_ind then Drule.asm_rl else
@@ -790,16 +789,16 @@
 
     (* abbrevs *)
 
-    val (_, ctxt1) = Variable.add_fixes (map (Name.name_of o fst o fst) cnames_syn) lthy;
+    val (_, ctxt1) = Variable.add_fixes (map (Binding.base_name o fst o fst) cnames_syn) lthy;
 
     fun get_abbrev ((name, atts), t) =
       if can (Logic.strip_assums_concl #> Logic.dest_equals) t then
         let
-          val _ = Name.name_of name = "" andalso null atts orelse
+          val _ = Binding.is_empty name andalso null atts orelse
             error "Abbreviations may not have names or attributes";
           val ((x, T), rhs) = LocalDefs.abs_def (snd (LocalDefs.cert_def ctxt1 t));
           val var =
-            (case find_first (fn ((c, _), _) => Name.name_of c = x) cnames_syn of
+            (case find_first (fn ((c, _), _) => Binding.base_name c = x) cnames_syn of
               NONE => error ("Undeclared head of abbreviation " ^ quote x)
             | SOME ((b, T'), mx) =>
                 if T <> T' then error ("Bad type specification for abbreviation " ^ quote x)
@@ -808,17 +807,17 @@
       else NONE;
 
     val abbrevs = map_filter get_abbrev spec;
-    val bs = map (Name.name_of o fst o fst) abbrevs;
+    val bs = map (Binding.base_name o fst o fst) abbrevs;
 
 
     (* predicates *)
 
     val pre_intros = filter_out (is_some o get_abbrev) spec;
-    val cnames_syn' = filter_out (member (op =) bs o Name.name_of o fst o fst) cnames_syn;
-    val cs = map (Free o apfst Name.name_of o fst) cnames_syn';
+    val cnames_syn' = filter_out (member (op =) bs o Binding.base_name o fst o fst) cnames_syn;
+    val cs = map (Free o apfst Binding.base_name o fst) cnames_syn';
     val ps = map Free pnames;
 
-    val (_, ctxt2) = lthy |> Variable.add_fixes (map (Name.name_of o fst o fst) cnames_syn');
+    val (_, ctxt2) = lthy |> Variable.add_fixes (map (Binding.base_name o fst o fst) cnames_syn');
     val _ = map (fn abbr => LocalDefs.fixed_abbrev abbr ctxt2) abbrevs;
     val ctxt3 = ctxt2 |> fold (snd oo LocalDefs.fixed_abbrev) abbrevs;
     val expand = Assumption.export_term ctxt3 lthy #> ProofContext.cert_term lthy;
@@ -846,11 +845,11 @@
     val intrs = map (apsnd the_single) specs;
     val monos = Attrib.eval_thms lthy raw_monos;
     val flags = {quiet_mode = false, verbose = verbose, kind = Thm.theoremK,
-      alt_name = Name.no_binding, coind = coind, no_elim = false, no_ind = false, skip_mono = false};
+      alt_name = Binding.empty, coind = coind, no_elim = false, no_ind = false, skip_mono = false};
   in
     lthy
     |> LocalTheory.set_group (serial_string ())
-    |> gen_add_inductive_i mk_def flags cs (map (apfst Name.name_of o fst) ps) intrs monos
+    |> gen_add_inductive_i mk_def flags cs (map (apfst Binding.base_name o fst) ps) intrs monos
   end;
 
 val add_inductive_i = gen_add_inductive_i add_ind_def;
@@ -858,7 +857,7 @@
 
 fun add_inductive_global group flags cnames_syn pnames pre_intros monos thy =
   let
-    val name = Sign.full_name thy (Name.name_of (fst (fst (hd cnames_syn))));
+    val name = Sign.full_name thy (fst (fst (hd cnames_syn)));
     val ctxt' = thy
       |> TheoryTarget.init NONE
       |> LocalTheory.set_group group
@@ -946,11 +945,11 @@
 fun flatten_specification specs = specs |> maps
   (fn (a, (concl, [])) => concl |> map
         (fn ((b, atts), [B]) =>
-              if Name.name_of a = "" then ((b, atts), B)
-              else if Name.name_of b = "" then ((a, atts), B)
+              if Binding.is_empty a then ((b, atts), B)
+              else if Binding.is_empty b then ((a, atts), B)
               else error "Illegal nested case names"
           | ((b, _), _) => error "Illegal simultaneous specification")
-    | (a, _) => error ("Illegal local specification parameters for " ^ quote (Name.name_of a)));
+    | (a, _) => error ("Illegal local specification parameters for " ^ quote (Binding.base_name a)));
 
 fun gen_ind_decl mk_def coind =
   P.fixes -- P.for_fixes --
--- a/src/HOL/Tools/inductive_realizer.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/inductive_realizer.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -338,7 +338,7 @@
             (Logic.strip_assums_concl rintr));
           val s' = Sign.base_name s;
           val T' = Logic.unvarifyT T
-        in (((Name.binding s', T'), NoSyn), (Const (s, T'), Free (s', T'))) end) rintrs));
+        in (((Binding.name s', T'), NoSyn), (Const (s, T'), Free (s', T'))) end) rintrs));
     val rlzparams = map (fn Var ((s, _), T) => (s, Logic.unvarifyT T))
       (List.take (snd (strip_comb
         (HOLogic.dest_Trueprop (Logic.strip_assums_concl (hd rintrs)))), nparms));
@@ -347,10 +347,10 @@
 
     val (ind_info, thy3') = thy2 |>
       InductivePackage.add_inductive_global (serial_string ())
-        {quiet_mode = false, verbose = false, kind = Thm.theoremK, alt_name = Name.no_binding,
+        {quiet_mode = false, verbose = false, kind = Thm.theoremK, alt_name = Binding.empty,
           coind = false, no_elim = false, no_ind = false, skip_mono = false}
         rlzpreds rlzparams (map (fn (rintr, intr) =>
-          ((Name.binding (Sign.base_name (name_of_thm intr)), []),
+          ((Binding.name (Sign.base_name (name_of_thm intr)), []),
            subst_atomic rlzpreds' (Logic.unvarify rintr)))
              (rintrs ~~ maps snd rss)) [] ||>
       Sign.absolute_path;
--- a/src/HOL/Tools/inductive_set_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/inductive_set_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/inductive_set_package.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Wrapper for defining inductive sets using package for inductive predicates,
@@ -13,13 +12,13 @@
   val pred_set_conv_att: attribute
   val add_inductive_i:
     InductivePackage.inductive_flags ->
-    ((Name.binding * typ) * mixfix) list ->
+    ((Binding.T * typ) * mixfix) list ->
     (string * typ) list ->
     (Attrib.binding * term) list -> thm list ->
     local_theory -> InductivePackage.inductive_result * local_theory
   val add_inductive: bool -> bool ->
-    (Name.binding * string option * mixfix) list ->
-    (Name.binding * string option * mixfix) list ->
+    (Binding.T * string option * mixfix) list ->
+    (Binding.T * string option * mixfix) list ->
     (Attrib.binding * string) list -> (Facts.ref * Attrib.src list) list ->
     local_theory -> InductivePackage.inductive_result * local_theory
   val codegen_preproc: theory -> thm list -> thm list
@@ -464,17 +463,17 @@
            | NONE => u)) |>
         Pattern.rewrite_term thy [] [to_pred_proc thy eqns'] |>
         eta_contract (member op = cs' orf is_pred pred_arities))) intros;
-    val cnames_syn' = map (fn (b, _) => (Name.map_name (suffix "p") b, NoSyn)) cnames_syn;
+    val cnames_syn' = map (fn (b, _) => (Binding.map_base (suffix "p") b, NoSyn)) cnames_syn;
     val monos' = map (to_pred [] (Context.Proof ctxt)) monos;
     val ({preds, intrs, elims, raw_induct, ...}, ctxt1) =
       InductivePackage.add_ind_def
-        {quiet_mode = quiet_mode, verbose = verbose, kind = kind, alt_name = Name.no_binding,
+        {quiet_mode = quiet_mode, verbose = verbose, kind = kind, alt_name = Binding.empty,
           coind = coind, no_elim = no_elim, no_ind = no_ind, skip_mono = skip_mono}
         cs' intros' monos' params1 cnames_syn' ctxt;
 
     (* define inductive sets using previously defined predicates *)
     val (defs, ctxt2) = fold_map (LocalTheory.define Thm.internalK)
-      (map (fn ((c_syn, (fs, U, _)), p) => (c_syn, (Attrib.no_binding,
+      (map (fn ((c_syn, (fs, U, _)), p) => (c_syn, (Attrib.empty_binding,
          fold_rev lambda params (HOLogic.Collect_const U $
            HOLogic.ap_split' fs U HOLogic.boolT (list_comb (p, params3))))))
          (cnames_syn ~~ cs_info ~~ preds)) ctxt1;
@@ -492,17 +491,17 @@
             (K (REPEAT (rtac ext 1) THEN simp_tac (HOL_basic_ss addsimps
               [def, mem_Collect_eq, split_conv]) 1))
         in
-          ctxt |> LocalTheory.note kind ((Name.binding (s ^ "p_" ^ s ^ "_eq"),
+          ctxt |> LocalTheory.note kind ((Binding.name (s ^ "p_" ^ s ^ "_eq"),
             [Attrib.internal (K pred_set_conv_att)]),
               [conv_thm]) |> snd
         end) (preds ~~ cs ~~ cs_info ~~ defs) ctxt2;
 
     (* convert theorems to set notation *)
     val rec_name =
-      if Binding.is_nothing alt_name then
-        Binding.binding (space_implode "_" (map (Name.name_of o fst) cnames_syn))
+      if Binding.is_empty alt_name then
+        Binding.name (space_implode "_" (map (Binding.base_name o fst) cnames_syn))
       else alt_name;
-    val cnames = map (Sign.full_name (ProofContext.theory_of ctxt3) o Name.name_of o #1) cnames_syn;  (* FIXME *)
+    val cnames = map (Sign.full_name (ProofContext.theory_of ctxt3) o #1) cnames_syn;  (* FIXME *)
     val (intr_names, intr_atts) = split_list (map fst intros);
     val raw_induct' = to_set [] (Context.Proof ctxt3) raw_induct;
     val (intrs', elims', induct, ctxt4) =
--- a/src/HOL/Tools/primrec_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/primrec_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/primrec_package.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen; Norbert Voelker, FernUni Hagen;
                 Florian Haftmann, TU Muenchen
 
@@ -8,12 +7,12 @@
 
 signature PRIMREC_PACKAGE =
 sig
-  val add_primrec: (Name.binding * typ option * mixfix) list ->
+  val add_primrec: (Binding.T * typ option * mixfix) list ->
     (Attrib.binding * term) list -> local_theory -> thm list * local_theory
-  val add_primrec_global: (Name.binding * typ option * mixfix) list ->
+  val add_primrec_global: (Binding.T * typ option * mixfix) list ->
     (Attrib.binding * term) list -> theory -> thm list * theory
   val add_primrec_overloaded: (string * (string * typ) * bool) list ->
-    (Name.binding * typ option * mixfix) list ->
+    (Binding.T * typ option * mixfix) list ->
     (Attrib.binding * term) list -> theory -> thm list * theory
 end;
 
@@ -195,8 +194,8 @@
     val def_name = Thm.def_name (Sign.base_name fname);
     val rhs = singleton (Syntax.check_terms ctxt) raw_rhs;
     val SOME var = get_first (fn ((b, _), mx) =>
-      if Name.name_of b = fname then SOME (b, mx) else NONE) fixes;
-  in (var, ((Name.binding def_name, []), rhs)) end;
+      if Binding.base_name b = fname then SOME (b, mx) else NONE) fixes;
+  in (var, ((Binding.name def_name, []), rhs)) end;
 
 
 (* find datatypes which contain all datatypes in tnames' *)
@@ -232,7 +231,7 @@
   let
     val (fixes, spec) = prepare_spec prep_spec lthy raw_fixes raw_spec;
     val eqns = fold_rev (process_eqn (fn v => Variable.is_fixed lthy v
-      orelse exists (fn ((w, _), _) => v = Name.name_of w) fixes) o snd) spec [];
+      orelse exists (fn ((w, _), _) => v = Binding.base_name w) fixes) o snd) spec [];
     val tnames = distinct (op =) (map (#1 o snd) eqns);
     val dts = find_dts (DatatypePackage.get_datatypes (ProofContext.theory_of lthy)) tnames tnames;
     val main_fns = map (fn (tname, {index, ...}) =>
@@ -248,7 +247,7 @@
     val _ = if gen_eq_set (op =) (names1, names2) then ()
       else primrec_error ("functions " ^ commas_quote names2 ^
         "\nare not mutually recursive");
-    val qualify = Name.qualified
+    val qualify = Binding.qualify
       (space_implode "_" (map (Sign.base_name o #1) defs));
     val spec' = (map o apfst o apfst) qualify spec;
     val simp_atts = map (Attrib.internal o K)
@@ -260,7 +259,7 @@
     |-> (fn defs => `(fn ctxt => prove_spec ctxt names1 rec_rewrites defs spec'))
     |-> (fn simps => fold_map (LocalTheory.note Thm.theoremK) simps)
     |-> (fn simps' => LocalTheory.note Thm.theoremK
-          ((qualify (Name.binding "simps"), simp_atts), maps snd simps'))
+          ((qualify (Binding.name "simps"), simp_atts), maps snd simps'))
     |>> snd
   end handle PrimrecError (msg, some_eqn) =>
     error ("Primrec definition error:\n" ^ msg ^ (case some_eqn
@@ -299,7 +298,7 @@
       P.name >> pair false) --| P.$$$ ")")) (false, "");
 
 val old_primrec_decl =
-  opt_unchecked_name -- Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Name.name_of) -- P.prop);
+  opt_unchecked_name -- Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.base_name) -- P.prop);
 
 fun pipe_error t = P.!!! (Scan.fail_with (K
   (cat_lines ["Equations must be separated by " ^ quote "|", quote t])));
--- a/src/HOL/Tools/recdef_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/recdef_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -268,8 +268,8 @@
       error ("No termination condition #" ^ string_of_int i ^
         " in recdef definition of " ^ quote name);
   in
-    Specification.theorem Thm.internalK NONE (K I) (Name.binding bname, atts)
-      [] (Element.Shows [(Attrib.no_binding, [(HOLogic.mk_Trueprop tc, [])])]) int lthy
+    Specification.theorem Thm.internalK NONE (K I) (Binding.name bname, atts)
+      [] (Element.Shows [(Attrib.empty_binding, [(HOLogic.mk_Trueprop tc, [])])]) int lthy
   end;
 
 val recdef_tc = gen_recdef_tc Attrib.intern_src Sign.intern_const;
@@ -299,7 +299,7 @@
 
 val recdef_decl =
   Scan.optional (P.$$$ "(" -- P.!!! (P.$$$ "permissive" -- P.$$$ ")") >> K false) true --
-  P.name -- P.term -- Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Name.name_of) -- P.prop)
+  P.name -- P.term -- Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.base_name) -- P.prop)
     -- Scan.option hints
   >> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map P.triple_swap eqs) src);
 
@@ -320,7 +320,7 @@
 val _ =
   OuterSyntax.local_theory_to_proof' "recdef_tc" "recommence proof of termination condition (TFL)"
     K.thy_goal
-    ((SpecParse.opt_thm_name ":" >> apfst Name.name_of) -- P.xname --
+    ((SpecParse.opt_thm_name ":" >> apfst Binding.base_name) -- P.xname --
         Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")")
       >> (fn ((thm_name, name), i) => recdef_tc thm_name name i));
 
--- a/src/HOL/Tools/record_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/record_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -346,19 +346,14 @@
 val get_updates = Symtab.lookup o #updates o get_sel_upd;
 fun get_simpset thy = Simplifier.theory_context thy (#simpset (get_sel_upd thy));
 
-fun put_sel_upd names simps thy =
-  let
-    val sels = map (rpair ()) names;
-    val upds = map (suffix updateN) names ~~ names;
-
-    val {records, sel_upd = {selectors, updates, simpset},
-      equalities, extinjects, extsplit, splits, extfields,fieldext} = RecordsData.get thy;
-    val data = make_record_data records
-      {selectors = Symtab.extend (selectors, sels),
-        updates = Symtab.extend (updates, upds),
-        simpset = Simplifier.addsimps (simpset, simps)}
-       equalities extinjects extsplit splits extfields fieldext;
-  in RecordsData.put data thy end;
+fun put_sel_upd names simps = RecordsData.map (fn {records,
+  sel_upd = {selectors, updates, simpset},
+    equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+  make_record_data records
+    {selectors = fold (fn name => Symtab.update (name, ())) names selectors,
+      updates = fold (fn name => Symtab.update ((suffix updateN) name, name)) names updates,
+      simpset = Simplifier.addsimps (simpset, simps)}
+      equalities extinjects extsplit splits extfields fieldext);
 
 
 (* access 'equalities' *)
@@ -1762,8 +1757,8 @@
     val external_names = NameSpace.external_names (Sign.naming_of thy);
 
     val alphas = map fst args;
-    val name = Sign.full_name thy bname;
-    val full = Sign.full_name_path thy bname;
+    val name = Sign.full_bname thy bname;
+    val full = Sign.full_bname_path thy bname;
     val base = Sign.base_name;
 
     val (bfields, field_syntax) = split_list (map (fn (x, T, mx) => ((x, T), mx)) raw_fields);
@@ -2253,7 +2248,7 @@
 
     (* errors *)
 
-    val name = Sign.full_name thy bname;
+    val name = Sign.full_bname thy bname;
     val err_dup_record =
       if is_none (get_record thy name) then []
       else ["Duplicate definition of record " ^ quote name];
--- a/src/HOL/Tools/refute.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/refute.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -313,18 +313,10 @@
 
   (* (string * string) -> theory -> theory *)
 
-  fun set_default_param (name, value) thy =
-  let
-    val {interpreters, printers, parameters} = RefuteData.get thy
-  in
-    RefuteData.put (case Symtab.lookup parameters name of
-      NONE   =>
+  fun set_default_param (name, value) = RefuteData.map 
+    (fn {interpreters, printers, parameters} =>
       {interpreters = interpreters, printers = printers,
-        parameters = Symtab.extend (parameters, [(name, value)])}
-    | SOME _ =>
-      {interpreters = interpreters, printers = printers,
-        parameters = Symtab.update (name, value) parameters}) thy
-  end;
+        parameters = Symtab.update (name, value) parameters});
 
 (* ------------------------------------------------------------------------- *)
 (* get_default_param: retrieves the value associated with 'name' from        *)
--- a/src/HOL/Tools/res_axioms.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/res_axioms.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -84,10 +84,10 @@
             val rhs = list_abs_free (map dest_Free args, HOLogic.choice_const T $ xtp)
                     (*Forms a lambda-abstraction over the formal parameters*)
             val (c, thy') =
-              Sign.declare_const [Markup.property_internal] ((Name.binding cname, cT), NoSyn) thy
+              Sign.declare_const [Markup.property_internal] ((Binding.name cname, cT), NoSyn) thy
             val cdef = cname ^ "_def"
             val thy'' = Theory.add_defs_i true false [(cdef, Logic.mk_equals (c, rhs))] thy'
-            val ax = Thm.axiom thy'' (Sign.full_name thy'' cdef)
+            val ax = Thm.axiom thy'' (Sign.full_bname thy'' cdef)
           in dec_sko (subst_bound (list_comb (c, args), p)) (ax :: axs, thy'') end
       | dec_sko (Const ("All", _) $ (xtp as Abs (a, T, p))) thx =
           (*Universal quant: insert a free variable into body and continue*)
--- a/src/HOL/Tools/specification_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/specification_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -233,7 +233,7 @@
 
 val specification_decl =
   P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
-          Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Name.name_of) -- P.prop)
+          Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.base_name) -- P.prop)
 
 val _ =
   OuterSyntax.command "specification" "define constants by specification" K.thy_goal
@@ -244,7 +244,7 @@
 val ax_specification_decl =
     P.name --
     (P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
-           Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Name.name_of) -- P.prop))
+           Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.base_name) -- P.prop))
 
 val _ =
   OuterSyntax.command "ax_specification" "define constants by specification" K.thy_goal
--- a/src/HOL/Tools/typecopy_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/typecopy_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -143,7 +143,7 @@
     |> TheoryTarget.instantiation ([tyco], vs, [HOLogic.class_eq])
     |> `(fn lthy => Syntax.check_term lthy def_eq)
     |-> (fn def_eq => Specification.definition
-         (NONE, (Attrib.no_binding, def_eq)))
+         (NONE, (Attrib.empty_binding, def_eq)))
     |-> (fn (_, (_, def_thm)) =>
        Class.prove_instantiation_exit_result Morphism.thm
          (fn _ => fn def_thm => Class.intro_classes_tac []
--- a/src/HOL/Tools/typedef_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Tools/typedef_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -80,7 +80,7 @@
   let
     val _ = Theory.requires thy "Typedef" "typedefs";
     val ctxt = ProofContext.init thy;
-    val full = Sign.full_name thy;
+    val full = Sign.full_bname thy;
 
     (*rhs*)
     val full_name = full name;
--- a/src/HOL/Typedef.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Typedef.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -145,7 +145,7 @@
     thy
     |> TheoryTarget.instantiation ([tyco], vs, @{sort itself})
     |> `(fn lthy => Syntax.check_term lthy eq)
-    |-> (fn eq => Specification.definition (NONE, (Attrib.no_binding, eq)))
+    |-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq)))
     |> snd
     |> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
     |> LocalTheory.exit_global
--- a/src/HOL/Typerep.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/Typerep.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -67,7 +67,7 @@
     thy
     |> TheoryTarget.instantiation ([tyco], vs, @{sort typerep})
     |> `(fn lthy => Syntax.check_term lthy eq)
-    |-> (fn eq => Specification.definition (NONE, (Attrib.no_binding, eq)))
+    |-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq)))
     |> snd
     |> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
     |> LocalTheory.exit_global
--- a/src/HOL/ex/Quickcheck.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/ex/Quickcheck.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -151,11 +151,11 @@
           thy
           |> TheoryTarget.instantiation ([tyco], vs, @{sort random})
           |> PrimrecPackage.add_primrec
-               [(Name.binding (fst (dest_Free random')), SOME (snd (dest_Free random')), NoSyn)]
-                 (map (fn eq => ((Name.no_binding, [del_func]), eq)) eqs')
+               [(Binding.name (fst (dest_Free random')), SOME (snd (dest_Free random')), NoSyn)]
+                 (map (fn eq => ((Binding.empty, [del_func]), eq)) eqs')
           |-> add_code
           |> `(fn lthy => Syntax.check_term lthy eq)
-          |-> (fn eq => Specification.definition (NONE, (Attrib.no_binding, eq)))
+          |-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq)))
           |> snd
           |> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
           |> LocalTheory.exit_global
--- a/src/HOL/ex/Term_Of_Syntax.thy	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOL/ex/Term_Of_Syntax.thy	Wed Dec 10 10:28:16 2008 +0100
@@ -10,7 +10,7 @@
 begin
 
 setup {*
-  Sign.declare_const [] ((Name.binding "rterm_of", @{typ "'a \<Rightarrow> 'b"}), NoSyn)
+  Sign.declare_const [] ((Binding.name "rterm_of", @{typ "'a \<Rightarrow> 'b"}), NoSyn)
   #> snd
 *}
 
--- a/src/HOLCF/Tools/domain/domain_axioms.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOLCF/Tools/domain/domain_axioms.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -120,7 +120,7 @@
 in (* local *)
 
 fun add_axioms (comp_dnam, eqs : eq list) thy' = let
-  val comp_dname = Sign.full_name thy' comp_dnam;
+  val comp_dname = Sign.full_bname thy' comp_dnam;
   val dnames = map (fst o fst) eqs;
   val x_name = idx_name dnames "x"; 
   fun copy_app dname = %%:(dname^"_copy")`Bound 0;
--- a/src/HOLCF/Tools/domain/domain_extender.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOLCF/Tools/domain/domain_extender.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -100,7 +100,7 @@
 fun gen_add_domain prep_typ (comp_dnam, eqs''') thy''' =
   let
     val dtnvs = map ((fn (dname,vs) => 
-			 (Sign.full_name thy''' dname, map (Syntax.read_typ_global thy''') vs))
+			 (Sign.full_bname thy''' dname, map (Syntax.read_typ_global thy''') vs))
                    o fst) eqs''';
     val cons''' = map snd eqs''';
     fun thy_type  (dname,tvars)  = (Sign.base_name dname, length tvars, NoSyn);
--- a/src/HOLCF/Tools/domain/domain_theorems.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOLCF/Tools/domain/domain_theorems.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -607,7 +607,7 @@
 
 val dnames = map (fst o fst) eqs;
 val conss  = map  snd        eqs;
-val comp_dname = Sign.full_name thy comp_dnam;
+val comp_dname = Sign.full_bname thy comp_dnam;
 
 val d = writeln("Proving induction properties of domain "^comp_dname^" ...");
 val pg = pg' thy;
--- a/src/HOLCF/Tools/fixrec_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOLCF/Tools/fixrec_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOLCF/Tools/fixrec_package.ML
-    ID:         $Id$
     Author:     Amber Telfer and Brian Huffman
 
 Recursive function definition package for HOLCF.
@@ -10,9 +9,9 @@
   val legacy_infer_term: theory -> term -> term
   val legacy_infer_prop: theory -> term -> term
   val add_fixrec: bool -> (Attrib.binding * string) list list -> theory -> theory
-  val add_fixrec_i: bool -> ((Name.binding * attribute list) * term) list list -> theory -> theory
+  val add_fixrec_i: bool -> ((Binding.T * attribute list) * term) list list -> theory -> theory
   val add_fixpat: Attrib.binding * string list -> theory -> theory
-  val add_fixpat_i: (Name.binding * attribute list) * term list -> theory -> theory
+  val add_fixpat_i: (Binding.T * attribute list) * term list -> theory -> theory
 end;
 
 structure FixrecPackage: FIXREC_PACKAGE =
@@ -227,7 +226,7 @@
     val lengths = map length blocks;
     
     val ((bindings, srcss), strings) = apfst split_list (split_list eqns);
-    val names = map Name.name_of bindings;
+    val names = map Binding.base_name bindings;
     val atts = map (map (prep_attrib thy)) srcss;
     val eqn_ts = map (prep_prop thy) strings;
     val rec_ts = map (fn eq => chead_of (fst (dest_eqs (Logic.strip_imp_concl eq)))
@@ -279,7 +278,7 @@
     val ts = map (prep_term thy) strings;
     val simps = map (fix_pat thy) ts;
   in
-    (snd o PureThy.add_thmss [((Name.name_of name, simps), atts)]) thy
+    (snd o PureThy.add_thmss [((Binding.base_name name, simps), atts)]) thy
   end;
 
 val add_fixpat = gen_add_fixpat Syntax.read_term_global Attrib.attribute;
--- a/src/HOLCF/Tools/pcpodef_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/HOLCF/Tools/pcpodef_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -54,7 +54,7 @@
 fun prepare_pcpodef prep_term pcpo def name (t, vs, mx) raw_set opt_morphs thy =
   let
     val ctxt = ProofContext.init thy;
-    val full = Sign.full_name thy;
+    val full = Sign.full_bname thy;
 
     (*rhs*)
     val full_name = full name;
@@ -94,7 +94,7 @@
           |> TheoryTarget.instantiation ([full_tname], lhs_tfrees, @{sort "Porder.po"});
         val less_def' = Syntax.check_term lthy3 less_def;
         val ((_, (_, less_definition')), lthy4) = lthy3
-          |> Specification.definition (NONE, ((Name.binding ("less_" ^ name ^ "_def"), []), less_def'));
+          |> Specification.definition (NONE, ((Binding.name ("less_" ^ name ^ "_def"), []), less_def'));
         val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy4);
         val less_definition = singleton (ProofContext.export lthy4 ctxt_thy) less_definition';
         val thy5 = lthy4
--- a/src/Pure/Concurrent/future.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Concurrent/future.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -30,21 +30,24 @@
   val enabled: unit -> bool
   type task = TaskQueue.task
   type group = TaskQueue.group
-  val thread_data: unit -> (string * task * group) option
-  type 'a T
-  val task_of: 'a T -> task
-  val group_of: 'a T -> group
-  val peek: 'a T -> 'a Exn.result option
-  val is_finished: 'a T -> bool
-  val future: group option -> task list -> bool -> (unit -> 'a) -> 'a T
-  val fork: (unit -> 'a) -> 'a T
-  val fork_background: (unit -> 'a) -> 'a T
-  val join_results: 'a T list -> 'a Exn.result list
-  val join_result: 'a T -> 'a Exn.result
-  val join: 'a T -> 'a
+  val thread_data: unit -> (string * task) option
+  type 'a future
+  val task_of: 'a future -> task
+  val group_of: 'a future -> group
+  val peek: 'a future -> 'a Exn.result option
+  val is_finished: 'a future -> bool
+  val value: 'a -> 'a future
+  val fork: (unit -> 'a) -> 'a future
+  val fork_group: group -> (unit -> 'a) -> 'a future
+  val fork_deps: 'b future list -> (unit -> 'a) -> 'a future
+  val fork_background: (unit -> 'a) -> 'a future
+  val join_results: 'a future list -> 'a Exn.result list
+  val join_result: 'a future -> 'a Exn.result
+  val join: 'a future -> 'a
+  val map: ('a -> 'b) -> 'a future -> 'b future
   val focus: task list -> unit
   val interrupt_task: string -> unit
-  val cancel: 'a T -> unit
+  val cancel: 'a future -> unit
   val shutdown: unit -> unit
 end;
 
@@ -63,7 +66,7 @@
 type task = TaskQueue.task;
 type group = TaskQueue.group;
 
-local val tag = Universal.tag () : (string * task * group) option Universal.tag in
+local val tag = Universal.tag () : (string * task) option Universal.tag in
   fun thread_data () = the_default NONE (Thread.getLocal tag);
   fun setmp_thread_data data f x = Library.setmp_thread_data tag (thread_data ()) (SOME data) f x;
 end;
@@ -71,7 +74,7 @@
 
 (* datatype future *)
 
-datatype 'a T = Future of
+datatype 'a future = Future of
  {task: task,
   group: group,
   result: 'a Exn.result option ref};
@@ -82,6 +85,11 @@
 fun peek (Future {result, ...}) = ! result;
 fun is_finished x = is_some (peek x);
 
+fun value x = Future
+ {task = TaskQueue.new_task (),
+  group = TaskQueue.new_group (),
+  result = ref (SOME (Exn.Result x))};
+
 
 
 (** scheduling **)
@@ -140,7 +148,7 @@
 fun execute name (task, group, run) =
   let
     val _ = trace_active ();
-    val ok = setmp_thread_data (name, task, group) run ();
+    val ok = setmp_thread_data (name, task) run ();
     val _ = SYNCHRONIZED "execute" (fn () =>
      (change queue (TaskQueue.finish task);
       if ok then ()
@@ -246,10 +254,10 @@
       change_result queue (TaskQueue.enqueue group deps pri run) before notify_all ());
   in Future {task = task, group = group, result = result} end;
 
-fun fork_common pri = future (Option.map #3 (thread_data ())) [] pri;
-
-fun fork e = fork_common true e;
-fun fork_background e = fork_common false e;
+fun fork e = future NONE [] true e;
+fun fork_group group e = future (SOME group) [] true e;
+fun fork_deps deps e = future NONE (map task_of deps) true e;
+fun fork_background e = future NONE [] false e;
 
 
 (* join: retrieve results *)
@@ -274,7 +282,7 @@
               (*alien thread -- refrain from contending for resources*)
               while exists (not o is_finished) xs
               do SYNCHRONIZED "join_thread" (fn () => wait "join_thread")
-          | SOME (name, task, _) =>
+          | SOME (name, task) =>
               (*proper task -- actively work towards results*)
               let
                 val unfinished = xs |> map_filter
@@ -292,6 +300,8 @@
 fun join_result x = singleton join_results x;
 fun join x = Exn.release (join_result x);
 
+fun map f x = fork_deps [x] (fn () => f (join x));
+
 
 (* misc operations *)
 
@@ -324,3 +334,6 @@
   else ();
 
 end;
+
+type 'a future = 'a Future.future;
+
--- a/src/Pure/Concurrent/par_list.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Concurrent/par_list.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -31,7 +31,7 @@
   if Future.enabled () then
     let
       val group = TaskQueue.new_group ();
-      val futures = map (fn x => Future.future (SOME group) [] true (fn () => f x)) xs;
+      val futures = map (fn x => Future.fork_group group (fn () => f x)) xs;
       val _ = List.app (ignore o Future.join_result) futures;
     in Future.join_results futures end
   else map (Exn.capture f) xs;
--- a/src/Pure/Concurrent/task_queue.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Concurrent/task_queue.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -8,6 +8,7 @@
 signature TASK_QUEUE =
 sig
   eqtype task
+  val new_task: unit -> task
   val str_of_task: task -> string
   eqtype group
   val new_group: unit -> group
@@ -34,8 +35,11 @@
 (* identifiers *)
 
 datatype task = Task of serial;
+fun new_task () = Task (serial ());
+
 fun str_of_task (Task i) = string_of_int i;
 
+
 datatype group = Group of serial * bool ref;
 
 fun new_group () = Group (serial (), ref true);
@@ -81,8 +85,7 @@
 
 fun enqueue (group as Group (gid, _)) deps pri job (Queue {groups, jobs, focus}) =
   let
-    val id = serial ();
-    val task = Task id;
+    val task as Task id = new_task ();
     val groups' = Inttab.cons_list (gid, task) groups;
     val jobs' = jobs
       |> IntGraph.new_node (id, (group, Job (pri, job))) |> fold (add_job task) deps;
--- a/src/Pure/General/binding.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/General/binding.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -15,17 +15,17 @@
 sig
   include BASIC_BINDING
   type T
-  val binding_pos: string * Position.T -> T
-  val binding: string -> T
-  val no_binding: T
-  val dest_binding: T -> (string * bool) list * string
-  val is_nothing: T -> bool
-  val pos_of: T -> Position.T
- 
-  val map_binding: ((string * bool) list * string -> (string * bool) list * string)
-    -> T -> T
+  val name_pos: string * Position.T -> T
+  val name: string -> T
+  val empty: T
+  val map_base: (string -> string) -> T -> T
+  val qualify: string -> T -> T
   val add_prefix: bool -> string -> T -> T
   val map_prefix: ((string * bool) list -> T -> T) -> T -> T
+  val is_empty: T -> bool
+  val base_name: T -> string
+  val pos_of: T -> Position.T
+  val dest: T -> (string * bool) list * string
   val display: T -> string
 end
 
@@ -44,22 +44,31 @@
 datatype T = Binding of ((string * bool) list * string) * Position.T;
   (* (prefix components (with mandatory flag), base name, position) *)
 
-fun binding_pos (name, pos) = Binding (([], name), pos);
-fun binding name = binding_pos (name, Position.none);
-val no_binding = binding "";
-
-fun pos_of (Binding (_, pos)) = pos;
-fun dest_binding (Binding (prefix_name, _)) = prefix_name;
+fun name_pos (name, pos) = Binding (([], name), pos);
+fun name name = name_pos (name, Position.none);
+val empty = name "";
 
 fun map_binding f (Binding (prefix_name, pos)) = Binding (f prefix_name, pos);
 
-fun is_nothing (Binding ((_, name), _)) = name = "";
+val map_base = map_binding o apsnd;
+
+fun qualify_base path name =
+  if path = "" orelse name = "" then name
+  else path ^ "." ^ name;
+
+val qualify = map_base o qualify_base;
+  (*FIXME should all operations on bare names move here from name_space.ML ?*)
 
 fun add_prefix sticky prfx b = if prfx = "" then error "empty name prefix"
   else (map_binding o apfst) (cons (prfx, sticky)) b;
 
 fun map_prefix f (Binding ((prefix, name), pos)) =
-  f prefix (binding_pos (name, pos));
+  f prefix (name_pos (name, pos));
+
+fun is_empty (Binding ((_, name), _)) = name = "";
+fun base_name (Binding ((_, name), _)) = name;
+fun pos_of (Binding (_, pos)) = pos;
+fun dest (Binding (prefix_name, _)) = prefix_name;
 
 fun display (Binding ((prefix, name), _)) =
   let
--- a/src/Pure/General/lazy.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/General/lazy.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -8,13 +8,13 @@
 
 signature LAZY =
 sig
-  type 'a T
-  val same: 'a T * 'a T -> bool
-  val lazy: (unit -> 'a) -> 'a T
-  val value: 'a -> 'a T
-  val peek: 'a T -> 'a Exn.result option
-  val force: 'a T -> 'a
-  val map_force: ('a -> 'a) -> 'a T -> 'a T
+  type 'a lazy
+  val same: 'a lazy * 'a lazy -> bool
+  val lazy: (unit -> 'a) -> 'a lazy
+  val value: 'a -> 'a lazy
+  val peek: 'a lazy -> 'a Exn.result option
+  val force: 'a lazy -> 'a
+  val map_force: ('a -> 'a) -> 'a lazy -> 'a lazy
 end
 
 structure Lazy :> LAZY =
@@ -22,13 +22,13 @@
 
 (* datatype *)
 
-datatype 'a lazy =
+datatype 'a T =
   Lazy of unit -> 'a |
   Result of 'a Exn.result;
 
-type 'a T = 'a lazy ref;
+type 'a lazy = 'a T ref;
 
-fun same (r1: 'a T, r2) = r1 = r2;
+fun same (r1: 'a lazy, r2) = r1 = r2;
 
 fun lazy e = ref (Lazy e);
 fun value x = ref (Result (Exn.Result x));
@@ -58,3 +58,6 @@
 fun map_force f = value o f o force;
 
 end;
+
+type 'a lazy = 'a Lazy.lazy;
+
--- a/src/Pure/General/name_space.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/General/name_space.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -30,25 +30,22 @@
   val get_accesses: T -> string -> xstring list
   val merge: T * T -> T
   type naming
-  val path_of: naming -> string
+  val default_naming: naming
+  val declare: naming -> Binding.T -> T -> string * T
+  val full_name: naming -> Binding.T -> string
   val external_names: naming -> string -> string list
-  val full: naming -> bstring -> string
-  val declare: naming -> string -> T -> T
-  val default_naming: naming
+  val path_of: naming -> string
   val add_path: string -> naming -> naming
   val no_base_names: naming -> naming
   val qualified_names: naming -> naming
   val sticky_prefix: string -> naming -> naming
-  val full_binding: naming -> Binding.T -> string
-  val declare_binding: naming -> Binding.T -> T -> string * T
   type 'a table = T * 'a Symtab.table
   val empty_table: 'a table
-  val table_declare: naming -> Binding.T * 'a
+  val bind: naming -> Binding.T * 'a
     -> 'a table -> string * 'a table (*exception Symtab.DUP*)
-  val table_declare_permissive: naming -> Binding.T * 'a
-    -> 'a table -> string * 'a table
-  val extend_table: naming -> (bstring * 'a) list -> 'a table -> 'a table
   val merge_tables: ('a * 'a -> bool) -> 'a table * 'a table -> 'a table
+  val join_tables: (string -> 'a * 'a -> 'a)
+    -> 'a table * 'a table -> 'a table
   val dest_table: 'a table -> (string * 'a) list
   val extern_table: 'a table -> (xstring * 'a) list
 end;
@@ -267,9 +264,9 @@
 
 (* declarations *)
 
-fun full (Naming (path, (qualify, _))) = qualify path;
+fun full_internal (Naming (path, (qualify, _))) = qualify path;
 
-fun declare naming name space =
+fun declare_internal naming name space =
   if is_hidden name then
     error ("Attempt to declare hidden name " ^ quote name)
   else
@@ -281,12 +278,16 @@
       val (accs, accs') = pairself (map implode_name) (accesses naming names);
     in space |> fold (add_name name) accs |> put_accesses name accs' end;
 
-fun declare_binding bnaming b =
+fun full_name naming b =
+  let val (prefix, bname) = Binding.dest b
+  in full_internal (apply_prefix prefix naming) bname end;
+
+fun declare bnaming b =
   let
-    val (prefix, bname) = Binding.dest_binding b;
+    val (prefix, bname) = Binding.dest b;
     val naming = apply_prefix prefix bnaming;
-    val name = full naming bname;
-  in declare naming name #> pair name end;
+    val name = full_internal naming bname;
+  in declare_internal naming name #> pair name end;
 
 
 
@@ -296,25 +297,17 @@
 
 val empty_table = (empty, Symtab.empty);
 
-fun gen_table_declare update naming (binding, x) (space, tab) =
+fun bind naming (b, x) (space, tab) =
   let
-    val (name, space') = declare_binding naming binding space;
-  in (name, (space', update (name, x) tab)) end;
-
-fun table_declare naming = gen_table_declare Symtab.update_new naming;
-fun table_declare_permissive naming = gen_table_declare Symtab.update naming;
-
-fun full_binding naming b =
-  let val (prefix, bname) = Binding.dest_binding b
-  in full (apply_prefix prefix naming) bname end;
-
-fun extend_table naming bentries (space, tab) =
-  let val entries = map (apfst (full naming)) bentries
-  in (fold (declare naming o #1) entries space, Symtab.extend (tab, entries)) end;
+    val (name, space') = declare naming b space;
+  in (name, (space', Symtab.update_new (name, x) tab)) end;
 
 fun merge_tables eq ((space1, tab1), (space2, tab2)) =
   (merge (space1, space2), Symtab.merge eq (tab1, tab2));
 
+fun join_tables f ((space1, tab1), (space2, tab2)) =
+  (merge (space1, space2), Symtab.join f (tab1, tab2));
+
 fun ext_table (space, tab) =
   Symtab.fold (fn (name, x) => cons ((name, extern space name), x)) tab []
   |> Library.sort_wrt (#2 o #1);
--- a/src/Pure/General/table.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/General/table.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -41,7 +41,6 @@
   val map_entry: key -> ('a -> 'a) -> 'a table -> 'a table
   val map_default: (key * 'a) -> ('a -> 'a) -> 'a table -> 'a table
   val make: (key * 'a) list -> 'a table                                (*exception DUP*)
-  val extend: 'a table * (key * 'a) list -> 'a table                   (*exception DUP*)
   val join: (key -> 'a * 'a -> 'a) (*exception DUP/SAME*) ->
     'a table * 'a table -> 'a table                                    (*exception DUP*)
   val merge: ('a * 'a -> bool) -> 'a table * 'a table -> 'a table      (*exception DUP*)
@@ -364,9 +363,7 @@
 
 (* simultaneous modifications *)
 
-fun extend (table, args) = fold update_new args table;
-
-fun make entries = extend (empty, entries);
+fun make entries = fold update_new entries empty;
 
 fun join f (table1, table2) =
   let fun add (key, y) tab = modify key (fn NONE => y | SOME x => f key (x, y)) tab;
--- a/src/Pure/Isar/args.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/args.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -35,7 +35,7 @@
   val name_source: T list -> string * T list
   val name_source_position: T list -> (SymbolPos.text * Position.T) * T list
   val name: T list -> string * T list
-  val binding: T list -> Name.binding * T list
+  val binding: T list -> Binding.T * T list
   val alt_name: T list -> string * T list
   val symbol: T list -> string * T list
   val liberal_name: T list -> string * T list
@@ -66,8 +66,8 @@
   val parse1: (string -> bool) -> OuterLex.token list -> T list * OuterLex.token list
   val attribs: (string -> string) -> T list -> src list * T list
   val opt_attribs: (string -> string) -> T list -> src list * T list
-  val thm_name: (string -> string) -> string -> T list -> (Name.binding * src list) * T list
-  val opt_thm_name: (string -> string) -> string -> T list -> (Name.binding * src list) * T list
+  val thm_name: (string -> string) -> string -> T list -> (Binding.T * src list) * T list
+  val opt_thm_name: (string -> string) -> string -> T list -> (Binding.T * src list) * T list
   val syntax: string -> ('b * T list -> 'a * ('b * T list)) -> src -> 'b -> 'a * 'b
   val context_syntax: string -> (Context.generic * T list -> 'a * (Context.generic * T list)) ->
     src -> Proof.context -> 'a * Proof.context
@@ -171,7 +171,7 @@
 val name_source_position = named >> T.source_position_of;
 
 val name = named >> T.content_of;
-val binding = P.position name >> Binding.binding_pos;
+val binding = P.position name >> Binding.name_pos;
 val alt_name = alt_string >> T.content_of;
 val symbol = symbolic >> T.content_of;
 val liberal_name = symbol || name;
@@ -280,8 +280,8 @@
 
 fun opt_thm_name intern s =
   Scan.optional
-    ((binding -- opt_attribs intern || attribs intern >> pair Name.no_binding) --| $$$ s)
-    (Name.no_binding, []);
+    ((binding -- opt_attribs intern || attribs intern >> pair Binding.empty) --| $$$ s)
+    (Binding.empty, []);
 
 
 
--- a/src/Pure/Isar/attrib.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/attrib.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/attrib.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Symbolic representation of attributes -- with name and syntax.
@@ -8,8 +7,8 @@
 signature ATTRIB =
 sig
   type src = Args.src
-  type binding = Name.binding * src list
-  val no_binding: binding
+  type binding = Binding.T * src list
+  val empty_binding: binding
   val print_attributes: theory -> unit
   val intern: theory -> xstring -> string
   val intern_src: theory -> src -> src
@@ -55,8 +54,8 @@
 
 type src = Args.src;
 
-type binding = Name.binding * src list;
-val no_binding: binding = (Name.no_binding, []);
+type binding = Binding.T * src list;
+val empty_binding: binding = (Binding.empty, []);
 
 
 
@@ -119,7 +118,7 @@
 fun attribute thy = attribute_i thy o intern_src thy;
 
 fun eval_thms ctxt args = ProofContext.note_thmss Thm.theoremK
-    [((Name.no_binding, []),
+    [((Binding.empty, []),
       map (apsnd (map (attribute (ProofContext.theory_of ctxt)))) args)] ctxt
   |> fst |> maps snd;
 
@@ -147,8 +146,8 @@
 fun add_attributes raw_attrs thy =
   let
     val new_attrs =
-      raw_attrs |> map (fn (name, att, comment) => (name, ((att, comment), stamp ())));
-    fun add attrs = NameSpace.extend_table (Sign.naming_of thy) new_attrs attrs
+      raw_attrs |> map (fn (name, att, comment) => (Binding.name name, ((att, comment), stamp ())));
+    fun add attrs = fold (snd oo NameSpace.bind (Sign.naming_of thy)) new_attrs attrs
       handle Symtab.DUP dup => error ("Duplicate declaration of attributes " ^ quote dup);
   in Attributes.map add thy end;
 
@@ -373,7 +372,7 @@
 fun register_config config thy =
   let
     val bname = Config.name_of config;
-    val name = Sign.full_name thy bname;
+    val name = Sign.full_bname thy bname;
   in
     thy
     |> add_attributes [(bname, syntax (Scan.lift (scan_config thy config) >> Morphism.form),
--- a/src/Pure/Isar/calculation.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/calculation.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -115,7 +115,7 @@
 
 fun print_calculation false _ _ = ()
   | print_calculation true ctxt calc = Pretty.writeln
-      (ProofContext.pretty_fact ctxt (ProofContext.full_name ctxt calculationN, calc));
+      (ProofContext.pretty_fact ctxt (ProofContext.full_bname ctxt calculationN, calc));
 
 
 (* also and finally *)
--- a/src/Pure/Isar/class.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/class.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -96,7 +96,7 @@
     thy
     |> fold2 add_constraint (map snd consts) no_constraints
     |> prove_interpretation tac (class_name_morphism class) (Locale.Locale class)
-          (map SOME inst, map (pair (Attrib.no_binding) o Thm.prop_of) def_facts)
+          (map SOME inst, map (pair (Attrib.empty_binding) o Thm.prop_of) def_facts)
     ||> fold2 add_constraint (map snd consts) constraints
   end;
 
@@ -476,7 +476,7 @@
     val inst = the_inst thy' class;
     val params = map (apsnd fst o snd) (these_params thy' [class]);
 
-    val c' = Sign.full_name thy' c;
+    val c' = Sign.full_bname thy' c;
     val dict' = Morphism.term phi dict;
     val dict_def = map_types Logic.unvarifyT dict';
     val ty' = Term.fastype_of dict_def;
@@ -485,7 +485,7 @@
     fun get_axiom thy = ((Thm.varifyT o Thm.axiom thy) c', thy);
   in
     thy'
-    |> Sign.declare_const pos ((Name.binding c, ty''), mx) |> snd
+    |> Sign.declare_const pos ((Binding.name c, ty''), mx) |> snd
     |> Thm.add_def false false (c, def_eq)
     |>> Thm.symmetric
     ||>> get_axiom
@@ -507,13 +507,13 @@
 
     val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty)))
       (these_operations thy [class]);
-    val c' = Sign.full_name thy' c;
+    val c' = Sign.full_bname thy' c;
     val rhs' = Pattern.rewrite_term thy unchecks [] rhs;
     val rhs'' = map_types Logic.varifyT rhs';
     val ty' = Term.fastype_of rhs';
   in
     thy'
-    |> Sign.add_abbrev (#1 prmode) pos (Name.binding c, map_types Type.strip_sorts rhs'') |> snd
+    |> Sign.add_abbrev (#1 prmode) pos (Binding.name c, map_types Type.strip_sorts rhs'') |> snd
     |> Sign.add_const_constraint (c', SOME ty')
     |> Sign.notation true prmode [(Const (c', ty'), mx)]
     |> not (#1 prmode = PrintMode.input) ? register_operation class (c', (rhs', NONE))
@@ -545,7 +545,7 @@
     val constrain = Element.Constrains ((map o apsnd o map_atyps)
       (K (TFree (Name.aT, base_sort))) supparams);
     fun fork_syn (Element.Fixes xs) =
-          fold_map (fn (c, ty, syn) => cons (Name.name_of c, syn) #> pair (c, ty, NoSyn)) xs
+          fold_map (fn (c, ty, syn) => cons (Binding.base_name c, syn) #> pair (c, ty, NoSyn)) xs
           #>> Element.Fixes
       | fork_syn x = pair x;
     fun fork_syntax elems =
@@ -574,14 +574,14 @@
     val raw_params = (snd o chop (length supparams)) all_params;
     fun add_const (v, raw_ty) thy =
       let
-        val c = Sign.full_name thy v;
+        val c = Sign.full_bname thy v;
         val ty = map_atyps (K (TFree (Name.aT, base_sort))) raw_ty;
         val ty0 = Type.strip_sorts ty;
         val ty' = map_atyps (K (TFree (Name.aT, [class]))) ty0;
         val syn = (the_default NoSyn o AList.lookup (op =) global_syntax) v;
       in
         thy
-        |> Sign.declare_const [] ((Name.binding v, ty0), syn)
+        |> Sign.declare_const [] ((Binding.name v, ty0), syn)
         |> snd
         |> pair ((v, ty), (c, ty'))
       end;
@@ -609,7 +609,7 @@
     |> add_consts bname class base_sort sups supparams global_syntax
     |-> (fn (param_map, params) => AxClass.define_class (bname, supsort)
           (map (fst o snd) params)
-          [((Name.binding (bname ^ "_" ^ AxClass.axiomsN), []), map (globalize param_map) raw_pred)]
+          [((Binding.name (bname ^ "_" ^ AxClass.axiomsN), []), map (globalize param_map) raw_pred)]
     #> snd
     #> `get_axiom
     #-> (fn assm_axiom => fold (Sign.add_const_constraint o apsnd SOME o snd) params
@@ -618,7 +618,7 @@
 
 fun gen_class prep_spec bname raw_supclasses raw_elems thy =
   let
-    val class = Sign.full_name thy bname;
+    val class = Sign.full_bname thy bname;
     val (((sups, supparams), (supsort, base_sort, mergeexpr)), (elems, global_syntax)) =
       prep_spec thy raw_supclasses raw_elems;
     val supconsts = map (apsnd fst o snd) (these_params thy sups);
--- a/src/Pure/Isar/code.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/code.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -15,7 +15,7 @@
   val add_default_eqn_attrib: Attrib.src
   val del_eqn: thm -> theory -> theory
   val del_eqns: string -> theory -> theory
-  val add_eqnl: string * (thm * bool) list Lazy.T -> theory -> theory
+  val add_eqnl: string * (thm * bool) list lazy -> theory -> theory
   val map_pre: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory
   val map_post: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory
   val add_inline: thm -> theory -> theory
@@ -114,7 +114,7 @@
 
 (* defining equations *)
 
-type eqns = bool * (thm * bool) list Lazy.T;
+type eqns = bool * (thm * bool) list lazy;
   (*default flag, theorems with linear flag (perhaps lazy)*)
 
 fun pretty_lthms ctxt r = case Lazy.peek r
--- a/src/Pure/Isar/constdefs.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/constdefs.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/constdefs.ML
-    ID:         $Id$
     Author:     Makarius, Hagia Maria Sion Abbey (Jerusalem)
 
 Old-style constant definitions, with type-inference and optional
@@ -9,12 +8,12 @@
 
 signature CONSTDEFS =
 sig
-  val add_constdefs: (Name.binding * string option) list *
-    ((Name.binding * string option * mixfix) option *
+  val add_constdefs: (Binding.T * string option) list *
+    ((Binding.T * string option * mixfix) option *
       (Attrib.binding * string)) list -> theory -> theory
-  val add_constdefs_i: (Name.binding * typ option) list *
-    ((Name.binding * typ option * mixfix) option *
-      ((Name.binding * attribute list) * term)) list -> theory -> theory
+  val add_constdefs_i: (Binding.T * typ option) list *
+    ((Binding.T * typ option * mixfix) option *
+      ((Binding.T * attribute list) * term)) list -> theory -> theory
 end;
 
 structure Constdefs: CONSTDEFS =
@@ -39,15 +38,15 @@
     val prop = prep_prop var_ctxt raw_prop;
     val (c, T) = #1 (LocalDefs.cert_def thy_ctxt (Logic.strip_imp_concl prop));
     val _ =
-      (case Option.map Name.name_of d of
+      (case Option.map Binding.base_name d of
         NONE => ()
       | SOME c' =>
           if c <> c' then
             err ("Head of definition " ^ quote c ^ " differs from declaration " ^ quote c') []
           else ());
 
-    val def = Term.subst_atomic [(Free (c, T), Const (Sign.full_name thy c, T))] prop;
-    val name = Thm.def_name_optional c (Name.name_of raw_name);
+    val def = Term.subst_atomic [(Free (c, T), Const (Sign.full_bname thy c, T))] prop;
+    val name = Thm.def_name_optional c (Binding.base_name raw_name);
     val atts = map (prep_att thy) raw_atts;
 
     val thy' =
--- a/src/Pure/Isar/element.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/element.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/element.ML
-    ID:         $Id$
     Author:     Makarius
 
 Explicit data structures for some Isar language elements, with derived
@@ -10,11 +9,11 @@
 sig
   datatype ('typ, 'term) stmt =
     Shows of (Attrib.binding * ('term * 'term list) list) list |
-    Obtains of (Name.binding * ((Name.binding * 'typ option) list * 'term list)) list
+    Obtains of (Binding.T * ((Binding.T * 'typ option) list * 'term list)) list
   type statement = (string, string) stmt
   type statement_i = (typ, term) stmt
   datatype ('typ, 'term, 'fact) ctxt =
-    Fixes of (Name.binding * 'typ option * mixfix) list |
+    Fixes of (Binding.T * 'typ option * mixfix) list |
     Constrains of (string * 'typ) list |
     Assumes of (Attrib.binding * ('term * 'term list) list) list |
     Defines of (Attrib.binding * ('term * 'term list)) list |
@@ -24,8 +23,8 @@
   val facts_map: (('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt) ->
    (Attrib.binding * ('fact * Attrib.src list) list) list ->
    (Attrib.binding * ('c * Attrib.src list) list) list
-  val map_ctxt: {name: Name.binding -> Name.binding,
-    var: Name.binding * mixfix -> Name.binding * mixfix,
+  val map_ctxt: {binding: Binding.T -> Binding.T,
+    var: Binding.T * mixfix -> Binding.T * mixfix,
     typ: 'typ -> 'a, term: 'term -> 'b, fact: 'fact -> 'c,
     attrib: Attrib.src -> Attrib.src} -> ('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt
   val map_ctxt_attrib: (Attrib.src -> Attrib.src) ->
@@ -56,7 +55,7 @@
   val rename_var_name: (string * (string * mixfix option)) list ->
     string * mixfix -> string * mixfix
   val rename_var: (string * (string * mixfix option)) list ->
-    Name.binding * mixfix -> Name.binding * mixfix
+    Binding.T * mixfix -> Binding.T * mixfix
   val rename_term: (string * (string * mixfix option)) list -> term -> term
   val rename_thm: (string * (string * mixfix option)) list -> thm -> thm
   val rename_morphism: (string * (string * mixfix option)) list -> morphism
@@ -76,9 +75,9 @@
     (Attrib.binding * (thm list * Attrib.src list) list) list ->
     (Attrib.binding * (thm list * Attrib.src list) list) list
   val activate: (typ, term, Facts.ref) ctxt list -> Proof.context ->
-    (context_i list * (Name.binding * Thm.thm list) list) * Proof.context
+    (context_i list * (Binding.T * Thm.thm list) list) * Proof.context
   val activate_i: context_i list -> Proof.context ->
-    (context_i list * (Name.binding * Thm.thm list) list) * Proof.context
+    (context_i list * (Binding.T * Thm.thm list) list) * Proof.context
 end;
 
 structure Element: ELEMENT =
@@ -90,7 +89,7 @@
 
 datatype ('typ, 'term) stmt =
   Shows of (Attrib.binding * ('term * 'term list) list) list |
-  Obtains of (Name.binding * ((Name.binding * 'typ option) list * 'term list)) list;
+  Obtains of (Binding.T * ((Binding.T * 'typ option) list * 'term list)) list;
 
 type statement = (string, string) stmt;
 type statement_i = (typ, term) stmt;
@@ -99,7 +98,7 @@
 (* context *)
 
 datatype ('typ, 'term, 'fact) ctxt =
-  Fixes of (Name.binding * 'typ option * mixfix) list |
+  Fixes of (Binding.T * 'typ option * mixfix) list |
   Constrains of (string * 'typ) list |
   Assumes of (Attrib.binding * ('term * 'term list) list) list |
   Defines of (Attrib.binding * ('term * 'term list)) list |
@@ -110,23 +109,23 @@
 
 fun facts_map f facts = Notes ("", facts) |> f |> (fn Notes (_, facts') => facts');
 
-fun map_ctxt {name, var, typ, term, fact, attrib} =
+fun map_ctxt {binding, var, typ, term, fact, attrib} =
   fn Fixes fixes => Fixes (fixes |> map (fn (x, T, mx) =>
        let val (x', mx') = var (x, mx) in (x', Option.map typ T, mx') end))
    | Constrains xs => Constrains (xs |> map (fn (x, T) =>
-       let val x' = Name.name_of (#1 (var (Name.binding x, NoSyn))) in (x', typ T) end))
+       let val x' = Binding.base_name (#1 (var (Binding.name x, NoSyn))) in (x', typ T) end))
    | Assumes asms => Assumes (asms |> map (fn ((a, atts), propps) =>
-      ((name a, map attrib atts), propps |> map (fn (t, ps) => (term t, map term ps)))))
+      ((binding a, map attrib atts), propps |> map (fn (t, ps) => (term t, map term ps)))))
    | Defines defs => Defines (defs |> map (fn ((a, atts), (t, ps)) =>
-      ((name a, map attrib atts), (term t, map term ps))))
+      ((binding a, map attrib atts), (term t, map term ps))))
    | Notes (kind, facts) => Notes (kind, facts |> map (fn ((a, atts), bs) =>
-      ((name a, map attrib atts), bs |> map (fn (ths, btts) => (fact ths, map attrib btts)))));
+      ((binding a, map attrib atts), bs |> map (fn (ths, btts) => (fact ths, map attrib btts)))));
 
 fun map_ctxt_attrib attrib =
-  map_ctxt {name = I, var = I, typ = I, term = I, fact = I, attrib = attrib};
+  map_ctxt {binding = I, var = I, typ = I, term = I, fact = I, attrib = attrib};
 
 fun morph_ctxt phi = map_ctxt
- {name = Morphism.name phi,
+ {binding = Morphism.binding phi,
   var = Morphism.var phi,
   typ = Morphism.typ phi,
   term = Morphism.term phi,
@@ -137,8 +136,8 @@
 (* logical content *)
 
 fun params_of (Fixes fixes) = fixes |> map
-    (fn (x, SOME T, _) => (Name.name_of x, T)
-      | (x, _, _) => raise TERM ("Untyped context element parameter " ^ quote (Name.display x), []))
+    (fn (x, SOME T, _) => (Binding.base_name x, T)
+      | (x, _, _) => raise TERM ("Untyped context element parameter " ^ quote (Binding.display x), []))
   | params_of _ = [];
 
 fun prems_of (Assumes asms) = maps (map fst o snd) asms
@@ -163,7 +162,7 @@
 
 fun pretty_name_atts ctxt (b, atts) sep =
   let
-    val name = Name.display b;
+    val name = Binding.display b;
   in if name = "" andalso null atts then []
     else [Pretty.block
       (Pretty.breaks (Pretty.str name :: Attrib.pretty_attribs ctxt atts @ [Pretty.str sep]))]
@@ -183,8 +182,8 @@
       Pretty.block (Pretty.breaks (prt_name_atts a ":" @ prt_terms (map fst ts)));
 
     fun prt_var (x, SOME T) = Pretty.block
-          [Pretty.str (Name.name_of x ^ " ::"), Pretty.brk 1, prt_typ T]
-      | prt_var (x, NONE) = Pretty.str (Name.name_of x);
+          [Pretty.str (Binding.base_name x ^ " ::"), Pretty.brk 1, prt_typ T]
+      | prt_var (x, NONE) = Pretty.str (Binding.base_name x);
     val prt_vars = separate (Pretty.keyword "and") o map prt_var;
 
     fun prt_obtain (_, ([], ts)) = Pretty.block (Pretty.breaks (prt_terms ts))
@@ -208,11 +207,11 @@
     fun prt_mixfix NoSyn = []
       | prt_mixfix mx = [Pretty.brk 2, Syntax.pretty_mixfix mx];
 
-    fun prt_fix (x, SOME T, mx) = Pretty.block (Pretty.str (Name.name_of x ^ " ::") ::
+    fun prt_fix (x, SOME T, mx) = Pretty.block (Pretty.str (Binding.base_name x ^ " ::") ::
           Pretty.brk 1 :: prt_typ T :: Pretty.brk 1 :: prt_mixfix mx)
-      | prt_fix (x, NONE, mx) = Pretty.block (Pretty.str (Name.name_of x) ::
+      | prt_fix (x, NONE, mx) = Pretty.block (Pretty.str (Binding.base_name x) ::
           Pretty.brk 1 :: prt_mixfix mx);
-    fun prt_constrain (x, T) = prt_fix (Name.binding x, SOME T, NoSyn);
+    fun prt_constrain (x, T) = prt_fix (Binding.name x, SOME T, NoSyn);
 
     fun prt_asm (a, ts) =
       Pretty.block (Pretty.breaks (prt_name_atts a ":" @ map (prt_term o fst) ts));
@@ -246,13 +245,13 @@
     else Pretty.command kind
   in Pretty.block (Pretty.fbreaks (head :: prts)) end;
 
-fun fix (x, T) = (Name.binding x, SOME T);
+fun fix (x, T) = (Binding.name x, SOME T);
 
 fun obtain prop ctxt =
   let
     val ((xs, prop'), ctxt') = Variable.focus prop ctxt;
     val As = Logic.strip_imp_prems (Thm.term_of prop');
-  in ((Name.no_binding, (map (fix o Term.dest_Free o Thm.term_of) xs, As)), ctxt') end;
+  in ((Binding.empty, (map (fix o Term.dest_Free o Thm.term_of) xs, As)), ctxt') end;
 
 in
 
@@ -275,9 +274,9 @@
     val (assumes, cases) = take_suffix (fn prem =>
       is_elim andalso concl aconv Logic.strip_assums_concl prem) prems;
   in
-    pretty_ctxt ctxt' (Fixes (map (fn (x, T) => (Name.binding x, SOME T, NoSyn)) fixes)) @
-    pretty_ctxt ctxt' (Assumes (map (fn t => (Attrib.no_binding, [(t, [])])) assumes)) @
-     (if null cases then pretty_stmt ctxt' (Shows [(Attrib.no_binding, [(concl, [])])])
+    pretty_ctxt ctxt' (Fixes (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) fixes)) @
+    pretty_ctxt ctxt' (Assumes (map (fn t => (Attrib.empty_binding, [(t, [])])) assumes)) @
+     (if null cases then pretty_stmt ctxt' (Shows [(Attrib.empty_binding, [(concl, [])])])
       else
         let val (clauses, ctxt'') = fold_map (obtain o cert) cases ctxt'
         in pretty_stmt ctxt'' (Obtains clauses) end)
@@ -396,9 +395,9 @@
 
 fun rename_var ren (b, mx) =
   let
-    val x = Name.name_of b;
+    val x = Binding.base_name b;
     val (x', mx') = rename_var_name ren (x, mx);
-  in (Name.binding x', mx') end;
+  in (Binding.name x', mx') end;
 
 fun rename_term ren (Free (x, T)) = Free (rename ren x, T)
   | rename_term ren (t $ u) = rename_term ren t $ rename_term ren u
@@ -419,7 +418,7 @@
   end;
 
 fun rename_morphism ren = Morphism.morphism
-  {name = I, var = rename_var ren, typ = I, term = rename_term ren, fact = map (rename_thm ren)};
+  {binding = I, var = rename_var ren, typ = I, term = rename_term ren, fact = map (rename_thm ren)};
 
 
 (* instantiate types *)
@@ -447,7 +446,7 @@
 fun instT_morphism thy env =
   let val thy_ref = Theory.check_thy thy in
     Morphism.morphism
-     {name = I, var = I,
+     {binding = I, var = I,
       typ = instT_type env,
       term = instT_term env,
       fact = map (fn th => instT_thm (Theory.deref thy_ref) env th)}
@@ -496,7 +495,7 @@
 fun inst_morphism thy envs =
   let val thy_ref = Theory.check_thy thy in
     Morphism.morphism
-     {name = I, var = I,
+     {binding = I, var = I,
       typ = instT_type (#1 envs),
       term = inst_term envs,
       fact = map (fn th => inst_thm (Theory.deref thy_ref) envs th)}
@@ -528,7 +527,7 @@
     val exp_term = Drule.term_rule thy (singleton exp_fact);
     val exp_typ = Logic.type_map exp_term;
     val morphism =
-      Morphism.morphism {name = I, var = I, typ = exp_typ, term = exp_term, fact = exp_fact};
+      Morphism.morphism {binding = I, var = I, typ = exp_typ, term = exp_term, fact = exp_fact};
   in facts_map (morph_ctxt morphism) facts end;
 
 
@@ -553,7 +552,7 @@
         val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs;
         val asms = defs' |> map (fn ((name, atts), (t, ps)) =>
             let val ((c, _), t') = LocalDefs.cert_def ctxt t
-            in (t', ((Name.map_name (Thm.def_name_optional c) name, atts), [(t', ps)])) end);
+            in (t', ((Binding.map_base (Thm.def_name_optional c) name, atts), [(t', ps)])) end);
         val (_, ctxt') =
           ctxt |> fold (Variable.auto_fixes o #1) asms
           |> ProofContext.add_assms_i LocalDefs.def_export (map #2 asms);
@@ -577,7 +576,7 @@
 
 fun prep_facts prep_name get intern ctxt elem = elem |> map_ctxt
      {var = I, typ = I, term = I,
-      name = Name.map_name prep_name,
+      binding = Binding.map_base prep_name,
       fact = get ctxt,
       attrib = intern (ProofContext.theory_of ctxt)};
 
--- a/src/Pure/Isar/expression.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/expression.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -8,8 +8,8 @@
 sig
   datatype 'term map = Positional of 'term option list | Named of (string * 'term) list;
   type 'term expr = (string * (string * 'term map)) list;
-  type expression = string expr * (Name.binding * string option * mixfix) list;
-  type expression_i = term expr * (Name.binding * typ option * mixfix) list;
+  type expression = string expr * (Binding.T * string option * mixfix) list;
+  type expression_i = term expr * (Binding.T * typ option * mixfix) list;
 
   (* Processing of context statements *)
   val read_statement: Element.context list -> (string * string list) list list ->
@@ -26,6 +26,10 @@
   (* Interpretation *)
   val sublocale_cmd: string -> expression -> theory -> Proof.state;
   val sublocale: string -> expression_i -> theory -> Proof.state;
+  val interpretation_cmd: expression -> theory -> Proof.state;
+  val interpretation: expression_i -> theory -> Proof.state;
+  val interpret_cmd: expression -> bool -> Proof.state -> Proof.state;
+  val interpret: expression_i -> bool -> Proof.state -> Proof.state;
 
   (* Debugging and development *)
   val parse_expression: OuterParse.token list -> expression * OuterParse.token list
@@ -47,8 +51,8 @@
 
 type 'term expr = (string * (string * 'term map)) list;
 
-type expression = string expr * (Name.binding * string option * mixfix) list;
-type expression_i = term expr * (Name.binding * typ option * mixfix) list;
+type expression = string expr * (Binding.T * string option * mixfix) list;
+type expression_i = term expr * (Binding.T * typ option * mixfix) list;
 
 
 (** Parsing and printing **)
@@ -128,8 +132,8 @@
         if null dups then () else error (message ^ commas dups)
       end;
 
-    fun match_bind (n, b) = (n = Name.name_of b);
-    fun bind_eq (b1, b2) = (Name.name_of b1 = Name.name_of b2);
+    fun match_bind (n, b) = (n = Binding.base_name b);
+    fun bind_eq (b1, b2) = (Binding.base_name b1 = Binding.base_name b2);
       (* FIXME: cannot compare bindings for equality. *)
 
     fun params_loc loc =
@@ -164,15 +168,15 @@
                   (* FIXME: should check for bindings being the same.
                      Instead we check for equal name and syntax. *)
                   if mx1 = mx2 then mx1
-                  else error ("Conflicting syntax for parameter" ^ quote (Name.display p) ^
+                  else error ("Conflicting syntax for parameter" ^ quote (Binding.display p) ^
                     " in expression.")) (ps, ps')
               in (i', ps'') end) is []
           in (ps', is') end;
 
     val (implicit, expr') = params_expr expr;
 
-    val implicit' = map (#1 #> Name.name_of) implicit;
-    val fixed' = map (#1 #> Name.name_of) fixed;
+    val implicit' = map (#1 #> Binding.base_name) implicit;
+    val fixed' = map (#1 #> Binding.base_name) fixed;
     val _ = reject_dups "Duplicate fixed parameter(s): " fixed';
     val implicit'' = if strict then []
       else let val _ = reject_dups
@@ -228,7 +232,7 @@
     val inst = Symtab.make insts'';
   in
     (Element.inst_morphism (ProofContext.theory_of ctxt) (instT, inst) $>
-      Morphism.name_morphism (Name.qualified prfx), ctxt')
+      Morphism.binding_morphism (Binding.qualify prfx), ctxt')
   end;
 
 
@@ -237,7 +241,7 @@
 (** Parsing **)
 
 fun parse_elem prep_typ prep_term ctxt elem =
-  Element.map_ctxt {name = I, var = I, typ = prep_typ ctxt,
+  Element.map_ctxt {binding = I, var = I, typ = prep_typ ctxt,
     term = prep_term ctxt, fact = I, attrib = I} elem;
 
 fun parse_concl prep_term ctxt concl =
@@ -316,7 +320,7 @@
       let val (vars, _) = prep_vars fixes ctxt
       in ctxt |> ProofContext.add_fixes_i vars |> snd end
   | declare_elem prep_vars (Constrains csts) ctxt =
-      ctxt |> prep_vars (map (fn (x, T) => (Name.binding x, SOME T, NoSyn)) csts) |> snd
+      ctxt |> prep_vars (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) csts) |> snd
   | declare_elem _ (Assumes _) ctxt = ctxt
   | declare_elem _ (Defines _) ctxt = ctxt
   | declare_elem _ (Notes _) ctxt = ctxt;
@@ -365,7 +369,8 @@
             val rev_frees =
               Term.fold_aterms (fn Free (x, T) =>
                 if AList.defined (op =) parms x then I else insert (op =) (x, T) | _ => I) t [];
-          in Term.list_all_free (rev rev_frees, t) end;
+          in Term.list_all_free (rev rev_frees, t) end; (* FIXME use fold Logic.all *)
+  (* FIXME consider closing in syntactic phase *)
 
         fun no_binds [] = []
           | no_binds _ = error "Illegal term bindings in context element";
@@ -373,13 +378,16 @@
         (case elem of
           Assumes asms => Assumes (asms |> map (fn (a, propps) =>
             (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps)))
-        | Defines defs => Defines (defs |> map (fn (a, (t, ps)) =>
-            (a, (close_frees (#2 (LocalDefs.cert_def ctxt t)), no_binds ps))))
+        | Defines defs => Defines (defs |> map (fn ((name, atts), (t, ps)) =>
+            let val ((c, _), t') = LocalDefs.cert_def ctxt (close_frees t)
+            in
+              ((Binding.map_base (Thm.def_name_optional c) name, atts), (t', no_binds ps))
+            end))
         | e => e)
       end;
 
 fun finish_primitive parms _ (Fixes fixes) = Fixes (map (fn (binding, _, mx) =>
-      let val x = Name.name_of binding
+      let val x = Binding.base_name binding
       in (binding, AList.lookup (op =) parms x, mx) end) fixes)
   | finish_primitive _ _ (Constrains _) = Constrains []
   | finish_primitive _ close (Assumes asms) = close (Assumes asms)
@@ -390,17 +398,17 @@
   let
     val thy = ProofContext.theory_of ctxt;
     val (parm_names, parm_types) = NewLocale.params_of thy loc |>
-      map (fn (b, SOME T, _) => (Name.name_of b, T)) |> split_list;
+      map (fn (b, SOME T, _) => (Binding.base_name b, T)) |> split_list;
     val (asm, defs) = NewLocale.specification_of thy loc;
     val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst) ctxt;
     val asm' = Option.map (Morphism.term morph) asm;
     val defs' = map (Morphism.term morph) defs;
     val text' = text |>
       (if is_some asm
-        then eval_text ctxt false (Assumes [(Attrib.no_binding, [(the asm', [])])])
+        then eval_text ctxt false (Assumes [(Attrib.empty_binding, [(the asm', [])])])
         else I) |>
       (if not (null defs)
-        then eval_text ctxt false (Defines (map (fn def => (Attrib.no_binding, (def, []))) defs'))
+        then eval_text ctxt false (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs'))
         else I)
 (* FIXME clone from new_locale.ML *)
   in ((loc, morph), text') end;
@@ -434,7 +442,7 @@
     fun prep_inst (loc, (prfx, inst)) (i, insts, ctxt) =
       let
         val (parm_names, parm_types) = NewLocale.params_of thy loc |>
-          map (fn (b, SOME T, _) => (Name.name_of b, T)) |> split_list;
+          map (fn (b, SOME T, _) => (Binding.base_name b, T)) |> split_list;
         val inst' = parse_inst parm_names inst ctxt;
         val parm_types' = map (TypeInfer.paramify_vars o
           Term.map_type_tvar (fn ((x, _), S) => TVar ((x, i), S)) o Logic.varifyT) parm_types;
@@ -462,12 +470,12 @@
 
     val fors = prep_vars fixed context |> fst;
     val ctxt = context |> ProofContext.add_fixes_i fors |> snd;
-    val (_, insts', ctxt') = fold prep_inst raw_insts (0, [], NewLocale.clear_idents ctxt);
+    val (_, insts', ctxt') = fold prep_inst raw_insts (0, [], NewLocale.clear_local_idents ctxt);
     val (_, elems'', ctxt'') = fold prep_elem raw_elems (insts', [], ctxt');
     val (insts, elems, concl, ctxt) = prep_concl raw_concl (insts', elems'', ctxt'');
 
     (* Retrieve parameter types *)
-    val xs = fold (fn Fixes fixes => (fn ps => ps @ map (Name.name_of o #1) fixes) |
+    val xs = fold (fn Fixes fixes => (fn ps => ps @ map (Binding.base_name o #1) fixes) |
       _ => fn ps => ps) (Fixes fors :: elems) [];
     val (Ts, ctxt') = fold_map ProofContext.inferred_param xs ctxt; 
     val parms = xs ~~ Ts;  (* params from expression and elements *)
@@ -532,16 +540,15 @@
 
 fun prep_declaration prep activate raw_import raw_elems context =
   let
-    val thy = ProofContext.theory_of context;
-
     val ((fixed, deps, elems, _), (parms, (spec, (_, _, defs)))) =
       prep false true context raw_import raw_elems [];
     (* Declare parameters and imported facts *)
     val context' = context |>
       ProofContext.add_fixes_i fixed |> snd |>
-      NewLocale.clear_idents |> fold (NewLocale.activate_facts thy) deps;
+      NewLocale.clear_local_idents |> fold NewLocale.activate_local_facts deps;
     val ((elems', _), _) = activate elems (ProofContext.set_stmt true context');
   in ((fixed, deps, elems'), (parms, spec, defs)) end;
+  (* FIXME check if defs used somewhere *)
 
 in
 
@@ -580,7 +587,7 @@
     val exp_term = Drule.term_rule thy (singleton exp_fact);
     val exp_typ = Logic.type_map exp_term;
     val export' =
-      Morphism.morphism {name = I, var = I, typ = exp_typ, term = exp_term, fact = exp_fact};
+      Morphism.morphism {binding = I, var = I, typ = exp_typ, term = exp_term, fact = exp_fact};
   in ((propss, deps, export'), goal_ctxt) end;
     
 in
@@ -634,7 +641,7 @@
 
 fun def_pred bname parms defs ts norm_ts thy =
   let
-    val name = Sign.full_name thy bname;
+    val name = Sign.full_bname thy bname;
 
     val (body, bodyT, body_eq) = atomize_spec thy norm_ts;
     val env = Term.add_term_free_names (body, []);
@@ -651,7 +658,7 @@
     val ([pred_def], defs_thy) =
       thy
       |> bodyT = propT ? Sign.add_advanced_trfuns ([], [], [aprop_tr' (length args) name], [])
-      |> Sign.declare_const [] ((Name.binding bname, predT), NoSyn) |> snd
+      |> Sign.declare_const [] ((Binding.name bname, predT), NoSyn) |> snd
       |> PureThy.add_defs false
         [((Thm.def_name bname, Logic.mk_equals (head, body)), [Thm.kind_internal])];
     val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head;
@@ -693,7 +700,7 @@
             |> Sign.add_path aname
             |> Sign.no_base_names
             |> PureThy.note_thmss Thm.internalK
-              [((Name.binding introN, []), [([intro], [NewLocale.unfold_attrib])])]
+              [((Binding.name introN, []), [([intro], [NewLocale.unfold_attrib])])]
             ||> Sign.restore_naming thy';
           in (SOME statement, SOME intro, axioms, thy'') end;
     val (b_pred, b_intro, b_axioms, thy'''') =
@@ -708,8 +715,8 @@
             |> Sign.add_path pname
             |> Sign.no_base_names
             |> PureThy.note_thmss Thm.internalK
-                 [((Name.binding introN, []), [([intro], [NewLocale.intro_attrib])]),
-                  ((Name.binding axiomsN, []),
+                 [((Binding.name introN, []), [([intro], [NewLocale.intro_attrib])]),
+                  ((Binding.name axiomsN, []),
                     [(map (Drule.standard o Element.conclude_witness) axioms, [])])]
             ||> Sign.restore_naming thy''';
         in (SOME statement, SOME intro, axioms, thy'''') end;
@@ -731,7 +738,7 @@
     let
       val defs' = map (fn (_, (def, _)) => def) defs
       val notes = map (fn (a, (def, _)) =>
-        (a, [([assume (cterm_of thy def)], [])])) defs
+        (a, [([Assumption.assume (cterm_of thy def)], [])])) defs
     in
       (Notes (Thm.definitionK, notes), defns @ defs')
     end
@@ -740,13 +747,12 @@
 fun gen_add_locale prep_decl
     bname predicate_name raw_imprt raw_body thy =
   let
-    val thy_ctxt = ProofContext.init thy;
-    val name = Sign.full_name thy bname;
+    val name = Sign.full_bname thy bname;
     val _ = NewLocale.test_locale thy name andalso
       error ("Duplicate definition of locale " ^ quote name);
 
-    val ((fixed, deps, body_elems), text as (parms, ((_, exts'), _), defs)) =
-      prep_decl raw_imprt raw_body thy_ctxt;
+    val ((fixed, deps, body_elems), text as (parms, ((_, exts'), _), _)) =
+      prep_decl raw_imprt raw_body (ProofContext.init thy);
     val ((a_statement, a_intro, a_axioms), (b_statement, b_intro, b_axioms), thy') =
       define_preds predicate_name text thy;
 
@@ -765,15 +771,15 @@
       fst |> map (Element.morph_ctxt satisfy) |>
       map_filter (fn Notes notes => SOME notes | _ => NONE) |>
       (if is_some asm
-        then cons (Thm.internalK, [((Name.binding (bname ^ "_" ^ axiomsN), []),
+        then cons (Thm.internalK, [((Binding.name (bname ^ "_" ^ axiomsN), []),
           [([assume (cterm_of thy' (the asm))], [(Attrib.internal o K) NewLocale.witness_attrib])])])
         else I);
 
     val deps' = map (fn (l, morph) => (l, morph $> satisfy)) deps;
 
     val loc_ctxt = thy' |>
-      NewLocale.register_locale name (extraTs, params)
-        (asm, map prop_of defs) ([], [])
+      NewLocale.register_locale bname (extraTs, params)
+        (asm, defns) ([], [])
         (map (fn n => (n, stamp ())) notes |> rev) (map (fn d => (d, stamp ())) deps' |> rev) |>
       NewLocale.init name
   in (name, loc_ctxt) end;
@@ -811,7 +817,13 @@
     fun store_dep ((name, morph), thms) =
       NewLocale.add_dependency target (name, morph $> Element.satisfy_morphism thms $> export);
 
-    fun after_qed results = fold store_dep (deps ~~ prep_result propss results);
+    fun after_qed results =
+      ProofContext.theory (
+        (* store dependencies *)
+        fold store_dep (deps ~~ prep_result propss results) #>
+        (* propagate registrations *)
+        (fn thy => fold_rev (fn reg => NewLocale.activate_global_facts reg)
+          (NewLocale.get_global_registrations thy) thy));
   in
     goal_ctxt |>
       Proof.theorem_i NONE after_qed (prep_propp propss) |>
@@ -826,5 +838,75 @@
 end;
 
 
+(** Interpretation in theories **)
+
+local
+
+fun gen_interpretation prep_expr
+    expression thy =
+  let
+    val ctxt = ProofContext.init thy;
+
+    val ((propss, regs, export), goal_ctxt) = prep_expr expression ctxt;
+    
+    fun store_reg ((name, morph), thms) =
+      let
+        val morph' = morph $> Element.satisfy_morphism thms $> export;
+      in
+        NewLocale.add_global_registration (name, morph') #>
+        NewLocale.activate_global_facts (name, morph')
+      end;
+
+    fun after_qed results =
+      ProofContext.theory (fold store_reg (regs ~~ prep_result propss results));
+  in
+    goal_ctxt |>
+      Proof.theorem_i NONE after_qed (prep_propp propss) |>
+      Element.refine_witness |> Seq.hd
+  end;
+
+in
+
+fun interpretation_cmd x = gen_interpretation read_goal_expression x;
+fun interpretation x = gen_interpretation cert_goal_expression x;
+
 end;
 
+
+(** Interpretation in proof contexts **)
+
+local
+
+fun gen_interpret prep_expr
+    expression int state =
+  let
+    val _ = Proof.assert_forward_or_chain state;
+    val ctxt = Proof.context_of state;
+
+    val ((propss, regs, export), goal_ctxt) = prep_expr expression ctxt;
+    
+    fun store_reg ((name, morph), thms) =
+      let
+        val morph' = morph $> Element.satisfy_morphism thms $> export;
+      in
+        NewLocale.activate_local_facts (name, morph')
+      end;
+
+    fun after_qed results =
+      Proof.map_context (fold store_reg (regs ~~ prep_result propss results)) #> Seq.single;
+  in
+    state |> Proof.map_context (K goal_ctxt) |>
+      Proof.local_goal (ProofDisplay.print_results int) (K I) ProofContext.bind_propp_i
+      "interpret" NONE after_qed (map (pair (Binding.empty, [])) (prep_propp propss)) |>
+      Element.refine_witness |> Seq.hd
+  end;
+
+in
+
+fun interpret_cmd x = gen_interpret read_goal_expression x;
+fun interpret x = gen_interpret cert_goal_expression x;
+
+end;
+
+end;
+
--- a/src/Pure/Isar/isar_cmd.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/isar_cmd.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -14,8 +14,8 @@
   val typed_print_translation: bool * (string * Position.T) -> theory -> theory
   val print_ast_translation: bool * (string * Position.T) -> theory -> theory
   val oracle: bstring -> SymbolPos.text * Position.T -> theory -> theory
-  val add_axioms: ((Name.binding * string) * Attrib.src list) list -> theory -> theory
-  val add_defs: (bool * bool) * ((Name.binding * string) * Attrib.src list) list -> theory -> theory
+  val add_axioms: ((Binding.T * string) * Attrib.src list) list -> theory -> theory
+  val add_defs: (bool * bool) * ((Binding.T * string) * Attrib.src list) list -> theory -> theory
   val declaration: string * Position.T -> local_theory -> local_theory
   val simproc_setup: string -> string list -> string * Position.T -> string list ->
     local_theory -> local_theory
@@ -167,7 +167,7 @@
 (* axioms *)
 
 fun add_axms f args thy =
-  f (map (fn ((b, ax), srcs) => ((Name.name_of b, ax), map (Attrib.attribute thy) srcs)) args) thy;
+  f (map (fn ((b, ax), srcs) => ((Binding.base_name b, ax), map (Attrib.attribute thy) srcs)) args) thy;
 
 val add_axioms = add_axms (snd oo PureThy.add_axioms_cmd);
 
--- a/src/Pure/Isar/isar_syn.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/isar_syn.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -272,7 +272,7 @@
 val _ =
   OuterSyntax.local_theory "declare" "declare theorems (improper)" K.thy_decl
     (P.and_list1 SpecParse.xthms1
-      >> (fn args => #2 o Specification.theorems_cmd "" [(Attrib.no_binding, flat args)]));
+      >> (fn args => #2 o Specification.theorems_cmd "" [(Attrib.empty_binding, flat args)]));
 
 
 (* name space entry path *)
@@ -396,7 +396,7 @@
           (begin ? Toplevel.print) o Toplevel.begin_local_theory begin
             (Locale.add_locale_cmd name expr elems #-> TheoryTarget.begin)));
 
-val opt_prefix = Scan.optional (P.binding --| P.$$$ ":") Name.no_binding;
+val opt_prefix = Scan.optional (P.binding --| P.$$$ ":") Binding.empty;
 
 val _ =
   OuterSyntax.command "sublocale"
@@ -413,7 +413,7 @@
       opt_prefix  -- SpecParse.locale_expr -- SpecParse.locale_insts
         >> (fn ((name, expr), insts) => Toplevel.print o
             Toplevel.theory_to_proof
-              (Locale.interpretation_cmd (Name.name_of name) expr insts)));
+              (Locale.interpretation_cmd (Binding.base_name name) expr insts)));
 
 val _ =
   OuterSyntax.command "interpret"
@@ -422,7 +422,7 @@
     (opt_prefix -- SpecParse.locale_expr -- SpecParse.locale_insts
       >> (fn ((name, expr), insts) => Toplevel.print o
           Toplevel.proof'
-            (fn int => Locale.interpret_cmd (Name.name_of name) expr insts int)));
+            (fn int => Locale.interpret_cmd (Binding.base_name name) expr insts int)));
 
 
 (* classes *)
@@ -483,7 +483,7 @@
 fun gen_theorem kind =
   OuterSyntax.local_theory_to_proof' kind ("state " ^ kind) K.thy_goal
     (Scan.optional (SpecParse.opt_thm_name ":" --|
-      Scan.ahead (SpecParse.locale_keyword || SpecParse.statement_keyword)) Attrib.no_binding --
+      Scan.ahead (SpecParse.locale_keyword || SpecParse.statement_keyword)) Attrib.empty_binding --
       SpecParse.general_statement >> (fn (a, (elems, concl)) =>
         (Specification.theorem_cmd kind NONE (K I) a elems concl)));
 
--- a/src/Pure/Isar/local_defs.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/local_defs.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -12,10 +12,10 @@
   val mk_def: Proof.context -> (string * term) list -> term list
   val expand: cterm list -> thm -> thm
   val def_export: Assumption.export
-  val add_defs: ((Name.binding * mixfix) * ((Name.binding * attribute list) * term)) list ->
+  val add_defs: ((Binding.T * mixfix) * ((Binding.T * attribute list) * term)) list ->
     Proof.context -> (term * (string * thm)) list * Proof.context
-  val add_def: (Name.binding * mixfix) * term -> Proof.context -> (term * thm) * Proof.context
-  val fixed_abbrev: (Name.binding * mixfix) * term -> Proof.context ->
+  val add_def: (Binding.T * mixfix) * term -> Proof.context -> (term * thm) * Proof.context
+  val fixed_abbrev: (Binding.T * mixfix) * term -> Proof.context ->
     (term * term) * Proof.context
   val export: Proof.context -> Proof.context -> thm -> thm list * thm
   val export_cterm: Proof.context -> Proof.context -> cterm -> cterm * thm
@@ -91,8 +91,8 @@
   let
     val ((bvars, mxs), specs) = defs |> split_list |>> split_list;
     val ((bfacts, atts), rhss) = specs |> split_list |>> split_list;
-    val xs = map Name.name_of bvars;
-    val names = map2 (Name.map_name o Thm.def_name_optional) xs bfacts;
+    val xs = map Binding.base_name bvars;
+    val names = map2 (Binding.map_base o Thm.def_name_optional) xs bfacts;
     val eqs = mk_def ctxt (xs ~~ rhss);
     val lhss = map (fst o Logic.dest_equals) eqs;
   in
@@ -105,7 +105,7 @@
   end;
 
 fun add_def (var, rhs) ctxt =
-  let val ([(lhs, (_, th))], ctxt') = add_defs [(var, ((Name.no_binding, []), rhs))] ctxt
+  let val ([(lhs, (_, th))], ctxt') = add_defs [(var, ((Binding.empty, []), rhs))] ctxt
   in ((lhs, th), ctxt') end;
 
 
--- a/src/Pure/Isar/local_theory.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/local_theory.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -19,16 +19,16 @@
   val raw_theory: (theory -> theory) -> local_theory -> local_theory
   val checkpoint: local_theory -> local_theory
   val full_naming: local_theory -> NameSpace.naming
-  val full_name: local_theory -> bstring -> string
+  val full_name: local_theory -> Binding.T -> string
   val theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
   val theory: (theory -> theory) -> local_theory -> local_theory
   val target_result: (Proof.context -> 'a * Proof.context) -> local_theory -> 'a * local_theory
   val target: (Proof.context -> Proof.context) -> local_theory -> local_theory
   val affirm: local_theory -> local_theory
   val pretty: local_theory -> Pretty.T list
-  val abbrev: Syntax.mode -> (Name.binding * mixfix) * term -> local_theory ->
+  val abbrev: Syntax.mode -> (Binding.T * mixfix) * term -> local_theory ->
     (term * term) * local_theory
-  val define: string -> (Name.binding * mixfix) * (Attrib.binding * term) -> local_theory ->
+  val define: string -> (Binding.T * mixfix) * (Attrib.binding * term) -> local_theory ->
     (term * (string * thm)) * local_theory
   val note: string -> Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory
   val notes: string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->
@@ -56,10 +56,10 @@
 
 type operations =
  {pretty: local_theory -> Pretty.T list,
-  abbrev: Syntax.mode -> (Name.binding * mixfix) * term -> local_theory ->
+  abbrev: Syntax.mode -> (Binding.T * mixfix) * term -> local_theory ->
     (term * term) * local_theory,
   define: string ->
-    (Name.binding * mixfix) * (Attrib.binding * term) -> local_theory ->
+    (Binding.T * mixfix) * (Attrib.binding * term) -> local_theory ->
     (term * (string * thm)) * local_theory,
   notes: string ->
     (Attrib.binding * (thm list * Attrib.src list) list) list ->
@@ -142,7 +142,7 @@
   |> NameSpace.sticky_prefix (#theory_prefix (get_lthy lthy))
   |> NameSpace.qualified_names;
 
-val full_name = NameSpace.full o full_naming;
+fun full_name naming = NameSpace.full_name (full_naming naming);
 
 fun theory_result f lthy = lthy |> raw_theory_result (fn thy => thy
   |> Sign.sticky_prefix (#theory_prefix (get_lthy lthy))
--- a/src/Pure/Isar/locale.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/locale.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -92,10 +92,10 @@
 
   (* Storing results *)
   val global_note_qualified: string ->
-    ((Name.binding * attribute list) * (thm list * attribute list) list) list ->
+    ((Binding.T * attribute list) * (thm list * attribute list) list) list ->
     theory -> (string * thm list) list * theory
   val local_note_qualified: string ->
-    ((Name.binding * attribute list) * (thm list * attribute list) list) list ->
+    ((Binding.T * attribute list) * (thm list * attribute list) list) list ->
     Proof.context -> (string * thm list) list * Proof.context
   val add_thmss: string -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->
     Proof.context -> Proof.context
@@ -104,11 +104,11 @@
   val add_declaration: string -> declaration -> Proof.context -> Proof.context
 
   (* Interpretation *)
-  val get_interpret_morph: theory -> (Name.binding -> Name.binding) -> string * string ->
+  val get_interpret_morph: theory -> (Binding.T -> Binding.T) -> string * string ->
     (Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) ->
     string -> term list -> Morphism.morphism
   val interpretation: (Proof.context -> Proof.context) ->
-    (Name.binding -> Name.binding) -> expr ->
+    (Binding.T -> Binding.T) -> expr ->
     term option list * (Attrib.binding * term) list ->
     theory ->
     (Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) * Proof.state
@@ -117,7 +117,7 @@
   val interpretation_in_locale: (Proof.context -> Proof.context) ->
     xstring * expr -> theory -> Proof.state
   val interpret: (Proof.state -> Proof.state Seq.seq) ->
-    (Name.binding -> Name.binding) -> expr ->
+    (Binding.T -> Binding.T) -> expr ->
     term option list * (Attrib.binding * term) list ->
     bool -> Proof.state ->
     (Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) * Proof.state
@@ -226,7 +226,7 @@
 (** management of registrations in theories and proof contexts **)
 
 type registration =
-  {prfx: (Name.binding -> Name.binding) * (string * string),
+  {prfx: (Binding.T -> Binding.T) * (string * string),
       (* first component: interpretation name morphism;
          second component: parameter prefix *)
     exp: Morphism.morphism,
@@ -248,18 +248,18 @@
     val join: T * T -> T
     val dest: theory -> T ->
       (term list *
-        (((Name.binding -> Name.binding) * (string * string)) *
+        (((Binding.T -> Binding.T) * (string * string)) *
          (Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) *
          Element.witness list *
          thm Termtab.table)) list
     val test: theory -> T * term list -> bool
     val lookup: theory ->
       T * (term list * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) ->
-      (((Name.binding -> Name.binding) * (string * string)) * Element.witness list * thm Termtab.table) option
-    val insert: theory -> term list -> ((Name.binding -> Name.binding) * (string * string)) ->
+      (((Binding.T -> Binding.T) * (string * string)) * Element.witness list * thm Termtab.table) option
+    val insert: theory -> term list -> ((Binding.T -> Binding.T) * (string * string)) ->
       (Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) ->
       T ->
-      T * (term list * (((Name.binding -> Name.binding) * (string * string)) * Element.witness list)) list
+      T * (term list * (((Binding.T -> Binding.T) * (string * string)) * Element.witness list)) list
     val add_witness: term list -> Element.witness -> T -> T
     val add_equation: term list -> thm -> T -> T
 (*
@@ -387,7 +387,7 @@
     (* 1st entry: locale namespace,
        2nd entry: locales of the theory *)
 
-  val empty = (NameSpace.empty, Symtab.empty);
+  val empty = NameSpace.empty_table;
   val copy = I;
   val extend = I;
 
@@ -403,8 +403,7 @@
       regs = merge_alists (op =) regs regs',
       intros = intros,
       dests = dests};
-  fun merge _ ((space1, locs1), (space2, locs2)) =
-    (NameSpace.merge (space1, space2), Symtab.join join_locales (locs1, locs2));
+  fun merge _ = NameSpace.join_tables join_locales;
 );
 
 
@@ -431,9 +430,9 @@
  of SOME loc => loc
   | NONE => error ("Unknown locale " ^ quote name);
 
-fun register_locale name loc thy =
-  thy |> LocalesData.map (fn (space, locs) =>
-    (Sign.declare_name thy name space, Symtab.update (name, loc) locs));
+fun register_locale bname loc thy =
+  thy |> LocalesData.map (NameSpace.bind (Sign.naming_of thy)
+    (Binding.name bname, loc) #> snd);
 
 fun change_locale name f thy =
   let
@@ -647,7 +646,7 @@
       | unify _ env = env;
     val (unifier, _) = fold unify (Ts' ~~ Us') (Vartab.empty, maxidx'');
     val Vs = map (Option.map (Envir.norm_type unifier)) Us';
-    val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (map_filter I Vs));
+    val unifier' = fold Vartab.update_new (frozen_tvars ctxt (map_filter I Vs)) unifier;
   in map (Option.map (Envir.norm_type unifier')) Vs end;
 
 fun params_of elemss =
@@ -712,7 +711,7 @@
       (Vartab.empty, maxidx);
 
     val parms' = map (apsnd (Envir.norm_type unifier)) (distinct (eq_fst (op =)) parms);
-    val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (map #2 parms'));
+    val unifier' = fold Vartab.update_new (frozen_tvars ctxt (map #2 parms')) unifier;
 
     fun inst_parms (i, ps) =
       List.foldr Term.add_typ_tfrees [] (map_filter snd ps)
@@ -811,7 +810,7 @@
 fun make_raw_params_elemss (params, tenv, syn) =
     [((("", map (fn p => (p, Symtab.lookup tenv p)) params), Assumed []),
       Int [Fixes (map (fn p =>
-        (Name.binding p, Symtab.lookup tenv p, Symtab.lookup syn p |> the)) params)])];
+        (Binding.name p, Symtab.lookup tenv p, Symtab.lookup syn p |> the)) params)])];
 
 
 (* flatten_expr:
@@ -954,7 +953,7 @@
           #> Binding.add_prefix false lprfx;
         val elem_morphism =
           Element.rename_morphism ren $>
-          Morphism.name_morphism add_prefices $>
+          Morphism.binding_morphism add_prefices $>
           Element.instT_morphism thy env;
         val elems' = map (Element.morph_ctxt elem_morphism) elems;
       in (((name, map (apsnd SOME) locale_params'), mode'), elems') end;
@@ -1003,7 +1002,7 @@
         val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs;
         val asms = defs' |> map (fn ((name, atts), (t, ps)) =>
             let val ((c, _), t') = LocalDefs.cert_def ctxt t
-            in (t', ((Name.map_name (Thm.def_name_optional c) name, atts), [(t', ps)])) end);
+            in (t', ((Binding.map_base (Thm.def_name_optional c) name, atts), [(t', ps)])) end);
         val (_, ctxt') =
           ctxt |> fold (Variable.auto_fixes o #1) asms
           |> ProofContext.add_assms_i LocalDefs.def_export (map #2 asms);
@@ -1101,15 +1100,15 @@
 *)
 
 fun flatten (ctxt, _) ((ids, syn), Elem (Fixes fixes)) = let
-        val ids' = ids @ [(("", map (Name.name_of o #1) fixes), ([], Assumed []))]
+        val ids' = ids @ [(("", map (Binding.base_name o #1) fixes), ([], Assumed []))]
       in
         ((ids',
          merge_syntax ctxt ids'
-           (syn, Symtab.make (map (fn fx => (Name.name_of (#1 fx), #3 fx)) fixes))
+           (syn, Symtab.make (map (fn fx => (Binding.base_name (#1 fx), #3 fx)) fixes))
            handle Symtab.DUP x => err_in_locale ctxt
              ("Conflicting syntax for parameter: " ^ quote x)
              (map #1 ids')),
-         [((("", map (rpair NONE o Name.name_of o #1) fixes), Assumed []), Ext (Fixes fixes))])
+         [((("", map (rpair NONE o Binding.base_name o #1) fixes), Assumed []), Ext (Fixes fixes))])
       end
   | flatten _ ((ids, syn), Elem elem) =
       ((ids @ [(("", []), ([], Assumed []))], syn), [((("", []), Assumed []), Ext elem)])
@@ -1129,7 +1128,7 @@
       let val (vars, _) = prep_vars fixes ctxt
       in ([], ctxt |> ProofContext.add_fixes_i vars |> snd) end
   | declare_ext_elem prep_vars (Constrains csts) ctxt =
-      let val (_, ctxt') = prep_vars (map (fn (x, T) => (Name.binding x, SOME T, NoSyn)) csts) ctxt
+      let val (_, ctxt') = prep_vars (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) csts) ctxt
       in ([], ctxt') end
   | declare_ext_elem _ (Assumes asms) ctxt = (map #2 asms, ctxt)
   | declare_ext_elem _ (Defines defs) ctxt = (map (fn (_, (t, ps)) => [(t, ps)]) defs, ctxt)
@@ -1253,7 +1252,7 @@
 
 
 fun finish_ext_elem parms _ (Fixes fixes, _) = Fixes (map (fn (b, _, mx) =>
-      let val x = Name.name_of b
+      let val x = Binding.base_name b
       in (b, AList.lookup (op =) parms x, mx) end) fixes)
   | finish_ext_elem parms _ (Constrains _, _) = Constrains []
   | finish_ext_elem _ close (Assumes asms, propp) =
@@ -1412,7 +1411,7 @@
       |> Element.morph_ctxt (Morphism.thm_morphism (Thm.transfer (ProofContext.theory_of ctxt)))
   | prep_facts prep_name get intern ctxt (Ext elem) = elem |> Element.map_ctxt
      {var = I, typ = I, term = I,
-      name = Name.map_name prep_name,
+      binding = Binding.map_base prep_name,
       fact = get ctxt,
       attrib = Args.assignable o intern (ProofContext.theory_of ctxt)};
 
@@ -1637,9 +1636,9 @@
 
 fun name_morph phi_name (lprfx, pprfx) b =
   b
-  |> (if not (Binding.is_nothing b) andalso pprfx <> ""
+  |> (if not (Binding.is_empty b) andalso pprfx <> ""
         then Binding.add_prefix false pprfx else I)
-  |> (if not (Binding.is_nothing b)
+  |> (if not (Binding.is_empty b)
         then Binding.add_prefix false lprfx else I)
   |> phi_name;
 
@@ -1651,9 +1650,9 @@
       (* FIXME sync with exp_fact *)
     val exp_typ = Logic.type_map exp_term;
     val export' =
-      Morphism.morphism {name = I, var = I, typ = exp_typ, term = exp_term, fact = exp_fact};
+      Morphism.morphism {binding = I, var = I, typ = exp_typ, term = exp_term, fact = exp_fact};
   in
-    Morphism.name_morphism (name_morph phi_name param_prfx) $>
+    Morphism.binding_morphism (name_morph phi_name param_prfx) $>
       Element.inst_morphism thy insts $>
       Element.satisfy_morphism prems $>
       Morphism.term_morphism (MetaSimplifier.rewrite_term thy eqns []) $>
@@ -1732,7 +1731,7 @@
     (fn (axiom, elems, params, decls, regs, intros, dests) =>
       (axiom, elems, params, add (decl, stamp ()) decls, regs, intros, dests))) #>
   add_thmss loc Thm.internalK
-    [((Name.no_binding, [Attrib.internal (decl_attrib decl)]), [([Drule.dummy_thm], [])])];
+    [((Binding.empty, [Attrib.internal (decl_attrib decl)]), [([Drule.dummy_thm], [])])];
 
 in
 
@@ -1789,7 +1788,7 @@
 
 fun def_pred bname parms defs ts norm_ts thy =
   let
-    val name = Sign.full_name thy bname;
+    val name = Sign.full_bname thy bname;
 
     val (body, bodyT, body_eq) = atomize_spec thy norm_ts;
     val env = Term.add_term_free_names (body, []);
@@ -1806,7 +1805,7 @@
     val ([pred_def], defs_thy) =
       thy
       |> bodyT = propT ? Sign.add_advanced_trfuns ([], [], [aprop_tr' (length args) name], [])
-      |> Sign.declare_const [] ((Name.binding bname, predT), NoSyn) |> snd
+      |> Sign.declare_const [] ((Binding.name bname, predT), NoSyn) |> snd
       |> PureThy.add_defs false
         [((Thm.def_name bname, Logic.mk_equals (head, body)), [Thm.kind_internal])];
     val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head;
@@ -1876,12 +1875,12 @@
             |> def_pred aname parms defs exts exts';
           val elemss' = change_assumes_elemss axioms elemss;
           val a_elem = [(("", []),
-            [Assumes [((Name.binding (pname ^ "_" ^ axiomsN), []), [(statement, [])])]])];
+            [Assumes [((Binding.name (pname ^ "_" ^ axiomsN), []), [(statement, [])])]])];
           val (_, thy'') =
             thy'
             |> Sign.add_path aname
             |> Sign.no_base_names
-            |> PureThy.note_thmss Thm.internalK [((Name.binding introN, []), [([intro], [])])]
+            |> PureThy.note_thmss Thm.internalK [((Binding.name introN, []), [([intro], [])])]
             ||> Sign.restore_naming thy';
         in ((elemss', [statement]), a_elem, [intro], thy'') end;
     val (predicate, stmt', elemss'', b_intro, thy'''') =
@@ -1894,14 +1893,14 @@
           val cstatement = Thm.cterm_of thy''' statement;
           val elemss'' = change_elemss_hyps axioms elemss';
           val b_elem = [(("", []),
-               [Assumes [((Name.binding (pname ^ "_" ^ axiomsN), []), [(statement, [])])]])];
+               [Assumes [((Binding.name (pname ^ "_" ^ axiomsN), []), [(statement, [])])]])];
           val (_, thy'''') =
             thy'''
             |> Sign.add_path pname
             |> Sign.no_base_names
             |> PureThy.note_thmss Thm.internalK
-                 [((Name.binding introN, []), [([intro], [])]),
-                  ((Name.binding axiomsN, []),
+                 [((Binding.name introN, []), [([intro], [])]),
+                  ((Binding.name axiomsN, []),
                     [(map (Drule.standard o Element.conclude_witness) axioms, [])])]
             ||> Sign.restore_naming thy''';
         in (([cstatement], axioms), [statement], elemss'' @ b_elem, [intro], thy'''') end;
@@ -1918,7 +1917,7 @@
 
 fun defines_to_notes is_ext thy (Defines defs) defns =
     let
-      val defs' = map (fn (_, (def, _)) => (Attrib.no_binding, (def, []))) defs
+      val defs' = map (fn (_, (def, _)) => (Attrib.empty_binding, (def, []))) defs
       val notes = map (fn (a, (def, _)) =>
         (a, [([assume (cterm_of thy def)], [])])) defs
     in
@@ -1940,7 +1939,7 @@
         "name" - locale with predicate named "name" *)
   let
     val thy_ctxt = ProofContext.init thy;
-    val name = Sign.full_name thy bname;
+    val name = Sign.full_bname thy bname;
     val _ = is_some (get_locale thy name) andalso
       error ("Duplicate definition of locale " ^ quote name);
 
@@ -1989,7 +1988,7 @@
       |> Sign.no_base_names
       |> PureThy.note_thmss Thm.assumptionK facts' |> snd
       |> Sign.restore_naming thy'
-      |> register_locale name {axiom = axs',
+      |> register_locale bname {axiom = axs',
         elems = map (fn e => (e, stamp ())) elems'',
         params = params_of elemss''' |> map (fn (x, SOME T) => ((x, T), the (Symtab.lookup syn x))),
         decls = ([], []),
@@ -2007,9 +2006,9 @@
 end;
 
 val _ = Context.>> (Context.map_theory
- (add_locale "" "var" empty [Fixes [(Name.binding (Name.internal "x"), NONE, NoSyn)]] #>
+ (add_locale "" "var" empty [Fixes [(Binding.name (Name.internal "x"), NONE, NoSyn)]] #>
   snd #> ProofContext.theory_of #>
-  add_locale "" "struct" empty [Fixes [(Name.binding (Name.internal "S"), NONE, Structure)]] #>
+  add_locale "" "struct" empty [Fixes [(Binding.name (Name.internal "S"), NONE, Structure)]] #>
   snd #> ProofContext.theory_of));
 
 
@@ -2378,7 +2377,7 @@
             fun activate_elem (loc, ps) (Notes (kind, facts)) thy =
                 let
                   val att_morphism =
-                    Morphism.name_morphism (name_morph phi_name param_prfx) $>
+                    Morphism.binding_morphism (name_morph phi_name param_prfx) $>
                     Morphism.thm_morphism satisfy $>
                     Element.inst_morphism thy insts $>
                     Morphism.thm_morphism disch;
@@ -2438,13 +2437,13 @@
   in
     state
     |> Proof.local_goal (ProofDisplay.print_results int) (K I) ProofContext.bind_propp_i
-      "interpret" NONE after_qed' (map (pair (Name.no_binding, [])) (prep_propp propss))
+      "interpret" NONE after_qed' (map (pair (Binding.empty, [])) (prep_propp propss))
     |> Element.refine_witness |> Seq.hd
     |> pair morphs
   end;
 
 fun standard_name_morph interp_prfx b =
-  if Binding.is_nothing b then b
+  if Binding.is_empty b then b
   else Binding.map_prefix (fn ((lprfx, _) :: pprfx) =>
     fold (Binding.add_prefix false o fst) pprfx
     #> interp_prfx <> "" ? Binding.add_prefix true interp_prfx
--- a/src/Pure/Isar/method.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/method.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -448,9 +448,9 @@
 fun add_methods raw_meths thy =
   let
     val new_meths = raw_meths |> map (fn (name, f, comment) =>
-      (name, ((f, comment), stamp ())));
+      (Binding.name name, ((f, comment), stamp ())));
 
-    fun add meths = NameSpace.extend_table (Sign.naming_of thy) new_meths meths
+    fun add meths = fold (snd oo NameSpace.bind (Sign.naming_of thy)) new_meths meths
       handle Symtab.DUP dup => error ("Duplicate declaration of method " ^ quote dup);
   in Methods.map add thy end;
 
--- a/src/Pure/Isar/new_locale.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/new_locale.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -8,10 +8,9 @@
 sig
   type locale
 
-  val clear_idents: Proof.context -> Proof.context
   val test_locale: theory -> string -> bool
-  val register_locale: string ->
-    (string * sort) list * (Name.binding * typ option * mixfix) list ->
+  val register_locale: bstring ->
+    (string * sort) list * (Binding.T * typ option * mixfix) list ->
     term option * term list ->
     (declaration * stamp) list * (declaration * stamp) list ->
     ((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * stamp) list ->
@@ -22,7 +21,8 @@
   val extern: theory -> string -> xstring
 
   (* Specification *)
-  val params_of: theory -> string -> (Name.binding * typ option * mixfix) list
+  val params_of: theory -> string -> (Binding.T * typ option * mixfix) list
+  val instance_of: theory -> string -> Morphism.morphism -> term list
   val specification_of: theory -> string -> term option * term list
   val declarations_of: theory -> string -> declaration list * declaration list
 
@@ -32,13 +32,13 @@
   val add_type_syntax: string -> declaration -> Proof.context -> Proof.context
   val add_term_syntax: string -> declaration -> Proof.context -> Proof.context
   val add_declaration: string -> declaration -> Proof.context -> Proof.context
-  val add_dependency: string -> (string * Morphism.morphism) -> Proof.context -> Proof.context
+  val add_dependency: string -> (string * Morphism.morphism) -> theory -> theory
 
   (* Activation *)
   val activate_declarations: theory -> string * Morphism.morphism ->
     Proof.context -> Proof.context
-  val activate_facts: theory -> string * Morphism.morphism ->
-    Proof.context -> Proof.context
+  val activate_global_facts: string * Morphism.morphism -> theory -> theory
+  val activate_local_facts: string * Morphism.morphism -> Proof.context -> Proof.context
 (*  val activate: string -> theory -> (Element.context_i -> 'a -> 'a) -> 'a -> 'a *)
   val init: string -> theory -> Proof.context
 
@@ -48,6 +48,11 @@
   val unfold_attrib: attribute
   val intro_locales_tac: bool -> Proof.context -> thm list -> tactic
 
+  (* Identifiers and registrations *)
+  val clear_local_idents: Proof.context -> Proof.context
+  val add_global_registration: (string * Morphism.morphism) -> theory -> theory
+  val get_global_registrations: theory -> (string * Morphism.morphism) list
+
   (* Diagnostic *)
   val print_locales: theory -> unit
   val print_locale: theory -> bool -> bstring -> unit
@@ -84,7 +89,7 @@
 
 datatype locale = Loc of {
   (* extensible lists are in reverse order: decls, notes, dependencies *)
-  parameters: (string * sort) list * (Name.binding * typ option * mixfix) list,
+  parameters: (string * sort) list * (Binding.T * typ option * mixfix) list,
     (* type and term parameters *)
   spec: term option * term list,
     (* assumptions (as a single predicate expression) and defines *)
@@ -104,7 +109,7 @@
   type T = NameSpace.T * locale Symtab.table;
     (* locale namespace and locales of the theory *)
 
-  val empty = (NameSpace.empty, Symtab.empty);
+  val empty = NameSpace.empty_table;
   val copy = I;
   val extend = I;
 
@@ -120,8 +125,7 @@
             dependencies = s_merge (dependencies, dependencies')
           }          
         end;
-  fun merge _ ((space1, locs1), (space2, locs2)) =
-    (NameSpace.merge (space1, space2), Symtab.join join_locales (locs1, locs2));
+  fun merge _ = NameSpace.join_tables join_locales;
 );
 
 val intern = NameSpace.intern o #1 o LocalesData.get;
@@ -136,11 +140,10 @@
 fun test_locale thy name = case get_locale thy name
  of SOME _ => true | NONE => false;
 
-fun register_locale name parameters spec decls notes dependencies thy =
-  thy |> LocalesData.map (fn (space, locs) =>
-    (Sign.declare_name thy name space, Symtab.update (name,
-      Loc {parameters = parameters, spec = spec, decls = decls, notes = notes,
-        dependencies = dependencies}) locs));
+fun register_locale bname parameters spec decls notes dependencies thy =
+  thy |> LocalesData.map (NameSpace.bind (Sign.naming_of thy) (Binding.name bname,
+    Loc {parameters = parameters, spec = spec, decls = decls, notes = notes,
+      dependencies = dependencies}) #> snd);
 
 fun change_locale name f thy =
   let
@@ -168,6 +171,10 @@
     val Loc {parameters = (_, params), ...} = the_locale thy name
   in params end;
 
+fun instance_of thy name morph =
+  params_of thy name |>
+    map ((fn (b, T, _) => Free (Binding.base_name b, the T)) #> Morphism.term morph);
+
 fun specification_of thy name =
   let
     val Loc {spec, ...} = the_locale thy name
@@ -193,19 +200,36 @@
 
 local
 
+datatype 'a delayed = Ready of 'a | ToDo of ('a delayed * 'a delayed);
+
 structure IdentifiersData = GenericDataFun
 (
-  type T = identifiers;
-  val empty = empty;
+  type T = identifiers delayed;
+  val empty = Ready empty;
   val extend = I;
-  fun merge _ = Library.merge (op =);  (* FIXME cannot do this properly without theory context *)
+  fun merge _ = ToDo;
 );
 
 in
 
-val get_idents = IdentifiersData.get o Context.Proof;
-val put_idents = Context.proof_map o IdentifiersData.put;
-val clear_idents = put_idents empty;
+fun finish thy (ToDo (i1, i2)) = merge (ident_eq thy) (finish thy i1, finish thy i2)
+  | finish _ (Ready ids) = ids;
+
+val _ = Context.>> (Context.map_theory (Theory.at_begin (fn thy =>
+  (case IdentifiersData.get (Context.Theory thy) of
+    Ready _ => NONE |
+    ids => SOME (Context.theory_map (IdentifiersData.put (Ready (finish thy ids))) thy))
+  )));
+  
+fun get_global_idents thy =
+  let val (Ready ids) = (IdentifiersData.get o Context.Theory) thy in ids end;
+val put_global_idents = Context.theory_map o IdentifiersData.put o Ready;
+val clear_global_idents = put_global_idents empty;
+
+fun get_local_idents ctxt =
+  let val (Ready ids) = (IdentifiersData.get o Context.Proof) ctxt in ids end;
+val put_local_idents = Context.proof_map o IdentifiersData.put o Ready;
+val clear_local_idents = put_local_idents empty;
 
 end;
 
@@ -222,8 +246,7 @@
   else
     let
       val Loc {parameters = (_, params), dependencies, ...} = the_locale thy name;
-      val instance = params |>
-        map ((fn (b, T, _) => Free (Name.name_of b, the T)) #> Morphism.term morph);
+      val instance = instance_of thy name morph;
     in
       if member (ident_eq thy) marked (name, instance)
       then (deps, marked)
@@ -242,9 +265,14 @@
 
 fun roundup thy activate_dep (name, morph) (marked, input) =
   let
-    val (dependencies', marked') = add thy 0 (name, morph) ([], marked);
+    (* Find all dependencies incuding new ones (which are dependencies enriching
+      existing registrations). *)
+    val (dependencies, marked') = add thy 0 (name, morph) ([], empty);
+    (* Filter out exisiting fragments. *)
+    val dependencies' = filter_out (fn (name, morph) =>
+      member (ident_eq thy) marked (name, instance_of thy name morph)) dependencies;
   in
-    (marked', input |> fold_rev (activate_dep thy) dependencies')
+    (merge (ident_eq thy) (marked, marked'), input |> fold_rev (activate_dep thy) dependencies')
   end;
 
 end;
@@ -273,15 +301,15 @@
 
 fun activate_all name thy activ_elem (marked, input) =
   let
-    val Loc {parameters = (_, params), spec = (asm, defs), dependencies, ...} =
+    val Loc {parameters = (_, params), spec = (asm, defs), ...} =
       the_locale thy name;
   in
     input |>
       (if not (null params) then activ_elem (Fixes params) else I) |>
       (* FIXME type parameters *)
-      (if is_some asm then activ_elem (Assumes [(Attrib.no_binding, [(the asm, [])])]) else I) |>
+      (if is_some asm then activ_elem (Assumes [(Attrib.empty_binding, [(the asm, [])])]) else I) |>
       (if not (null defs)
-        then activ_elem (Defines (map (fn def => (Attrib.no_binding, (def, []))) defs))
+        then activ_elem (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs))
         else I) |>
       pair marked |> roundup thy (activate_notes activ_elem) (name, Morphism.identity)
   end;
@@ -291,16 +319,21 @@
 
 local
 
-fun init_elem (Fixes fixes) ctxt = ctxt |>
+fun init_global_elem (Notes (kind, facts)) thy =
+      let
+        val facts' = Attrib.map_facts (Attrib.attribute_i thy) facts
+      in Locale.global_note_qualified kind facts' thy |> snd end
+
+fun init_local_elem (Fixes fixes) ctxt = ctxt |>
       ProofContext.add_fixes_i fixes |> snd
-  | init_elem (Assumes assms) ctxt =
+  | init_local_elem (Assumes assms) ctxt =
       let
         val assms' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) assms
       in
         ctxt |> fold Variable.auto_fixes (maps (map fst o snd) assms') |>
           ProofContext.add_assms_i Assumption.assume_export assms' |> snd
      end 
-  | init_elem (Defines defs) ctxt =
+  | init_local_elem (Defines defs) ctxt =
       let
         val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs
       in
@@ -308,7 +341,7 @@
           ProofContext.add_assms_i LocalDefs.def_export (map (fn (attn, t) => (attn, [t])) defs') |>
           snd
       end
-  | init_elem (Notes (kind, facts)) ctxt =
+  | init_local_elem (Notes (kind, facts)) ctxt =
       let
         val facts' = Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of ctxt)) facts
       in Locale.local_note_qualified kind facts' ctxt |> snd end
@@ -319,13 +352,19 @@
 in
 
 fun activate_declarations thy dep ctxt =
-  roundup thy activate_decls dep (get_idents ctxt, ctxt) |> uncurry put_idents;
-  
-fun activate_facts thy dep ctxt =
-  roundup thy (activate_notes init_elem) dep (get_idents ctxt, ctxt) |> uncurry put_idents;
+  roundup thy activate_decls dep (get_local_idents ctxt, ctxt) |> uncurry put_local_idents;
+
+fun activate_global_facts dep thy =
+  roundup thy (activate_notes init_global_elem) dep (get_global_idents thy, thy) |>
+  uncurry put_global_idents;
 
-fun init name thy = activate_all name thy init_elem (empty, ProofContext.init thy) |>
-  uncurry put_idents;
+fun activate_local_facts dep ctxt =
+  roundup (ProofContext.theory_of ctxt) (activate_notes init_local_elem) dep
+    (get_local_idents ctxt, ctxt) |>
+  uncurry put_local_idents;
+
+fun init name thy = activate_all name thy init_local_elem (empty, ProofContext.init thy) |>
+  uncurry put_local_idents;
 
 fun print_locale thy show_facts name =
   let
@@ -366,7 +405,7 @@
     (fn (parameters, spec, decls, notes, dependencies) =>
       (parameters, spec, add (decl, stamp ()) decls, notes, dependencies))) #>
   add_thmss loc Thm.internalK
-    [((Name.no_binding, [Attrib.internal (decl_attrib decl)]), [([Drule.dummy_thm], [])])];
+    [((Binding.empty, [Attrib.internal (decl_attrib decl)]), [([Drule.dummy_thm], [])])];
 
 in
 
@@ -379,9 +418,9 @@
 (* Dependencies *)
 
 fun add_dependency loc dep =
-  ProofContext.theory (change_locale loc
+  change_locale loc
     (fn (parameters, spec, decls, notes, dependencies) =>
-      (parameters, spec, decls, notes, (dep, stamp ()) :: dependencies)));
+      (parameters, spec, decls, notes, (dep, stamp ()) :: dependencies));
 
 
 (*** Reasoning about locales ***)
@@ -414,5 +453,22 @@
       "back-chain all introduction rules of locales")]));
 
 
+(*** Registrations: interpretations in theories ***)
+
+(* FIXME only global variant needed *)
+structure RegistrationsData = GenericDataFun
+(
+  type T = ((string * Morphism.morphism) * stamp) list;
+    (* registrations, in reverse order of declaration *)
+  val empty = [];
+  val extend = I;
+  fun merge _ = Library.merge (eq_snd (op =));
+    (* FIXME consolidate with dependencies, consider one data slot only *)
+);
+
+val get_global_registrations = map fst o RegistrationsData.get o Context.Theory;
+fun add_global_registration reg =
+  (Context.theory_map o RegistrationsData.map) (cons (reg, stamp ()));
+
 end;
 
--- a/src/Pure/Isar/object_logic.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/object_logic.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -89,7 +89,7 @@
 fun typedecl (raw_name, vs, mx) thy =
   let
     val base_sort = get_base_sort thy;
-    val name = Sign.full_name thy (Syntax.type_name raw_name mx);
+    val name = Sign.full_bname thy (Syntax.type_name raw_name mx);
     val _ = has_duplicates (op =) vs andalso
       error ("Duplicate parameters in type declaration: " ^ quote name);
     val n = length vs;
@@ -107,7 +107,7 @@
 local
 
 fun gen_add_judgment add_consts (bname, T, mx) thy =
-  let val c = Sign.full_name thy (Syntax.const_name bname mx) in
+  let val c = Sign.full_bname thy (Syntax.const_name bname mx) in
     thy
     |> add_consts [(bname, T, mx)]
     |> (fn thy' => Theory.add_deps c (c, Sign.the_const_type thy' c) [] thy')
--- a/src/Pure/Isar/obtain.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/obtain.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -40,16 +40,16 @@
 signature OBTAIN =
 sig
   val thatN: string
-  val obtain: string -> (Name.binding * string option * mixfix) list ->
+  val obtain: string -> (Binding.T * string option * mixfix) list ->
     (Attrib.binding * (string * string list) list) list ->
     bool -> Proof.state -> Proof.state
-  val obtain_i: string -> (Name.binding * typ option * mixfix) list ->
-    ((Name.binding * attribute list) * (term * term list) list) list ->
+  val obtain_i: string -> (Binding.T * typ option * mixfix) list ->
+    ((Binding.T * attribute list) * (term * term list) list) list ->
     bool -> Proof.state -> Proof.state
   val result: (Proof.context -> tactic) -> thm list -> Proof.context ->
     (cterm list * thm list) * Proof.context
-  val guess: (Name.binding * string option * mixfix) list -> bool -> Proof.state -> Proof.state
-  val guess_i: (Name.binding * typ option * mixfix) list -> bool -> Proof.state -> Proof.state
+  val guess: (Binding.T * string option * mixfix) list -> bool -> Proof.state -> Proof.state
+  val guess_i: (Binding.T * typ option * mixfix) list -> bool -> Proof.state -> Proof.state
 end;
 
 structure Obtain: OBTAIN =
@@ -122,7 +122,7 @@
     (*obtain vars*)
     val (vars, vars_ctxt) = prep_vars raw_vars ctxt;
     val (_, fix_ctxt) = vars_ctxt |> ProofContext.add_fixes_i vars;
-    val xs = map (Name.name_of o #1) vars;
+    val xs = map (Binding.base_name o #1) vars;
 
     (*obtain asms*)
     val (asms_ctxt, proppss) = prep_propp (fix_ctxt, map snd raw_asms);
@@ -156,14 +156,14 @@
   in
     state
     |> Proof.enter_forward
-    |> Proof.have_i NONE (K Seq.single) [((Name.no_binding, []), [(obtain_prop, [])])] int
+    |> Proof.have_i NONE (K Seq.single) [((Binding.empty, []), [(obtain_prop, [])])] int
     |> Proof.proof (SOME Method.succeed_text) |> Seq.hd
-    |> Proof.fix_i [(Name.binding thesisN, NONE, NoSyn)]
+    |> Proof.fix_i [(Binding.name thesisN, NONE, NoSyn)]
     |> Proof.assume_i
-      [((Name.binding that_name, [ContextRules.intro_query NONE]), [(that_prop, [])])]
+      [((Binding.name that_name, [ContextRules.intro_query NONE]), [(that_prop, [])])]
     |> `Proof.the_facts
     ||> Proof.chain_facts chain_facts
-    ||> Proof.show_i NONE after_qed [((Name.no_binding, []), [(thesis, [])])] false
+    ||> Proof.show_i NONE after_qed [((Binding.empty, []), [(thesis, [])])] false
     |-> Proof.refine_insert
   end;
 
@@ -261,7 +261,7 @@
 
 fun inferred_type (binding, _, mx) ctxt =
   let
-    val x = Name.name_of binding;
+    val x = Binding.base_name binding;
     val (T, ctxt') = ProofContext.inferred_param x ctxt
   in ((x, T, mx), ctxt') end;
 
@@ -294,9 +294,9 @@
       in
         state'
         |> Proof.map_context (K ctxt')
-        |> Proof.fix_i (map (fn ((x, T), mx) => (Name.binding x, SOME T, mx)) parms)
+        |> Proof.fix_i (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) parms)
         |> `Proof.context_of |-> (fn fix_ctxt => Proof.assm_i
-          (obtain_export fix_ctxt rule (map cert ts)) [((Name.no_binding, []), asms)])
+          (obtain_export fix_ctxt rule (map cert ts)) [((Binding.empty, []), asms)])
         |> Proof.add_binds_i AutoBind.no_facts
       end;
 
@@ -311,10 +311,10 @@
     state
     |> Proof.enter_forward
     |> Proof.begin_block
-    |> Proof.fix_i [(Name.binding AutoBind.thesisN, NONE, NoSyn)]
+    |> Proof.fix_i [(Binding.name AutoBind.thesisN, NONE, NoSyn)]
     |> Proof.chain_facts chain_facts
     |> Proof.local_goal print_result (K I) (apsnd (rpair I))
-      "guess" before_qed after_qed [((Name.no_binding, []), [Logic.mk_term goal, goal])]
+      "guess" before_qed after_qed [((Binding.empty, []), [Logic.mk_term goal, goal])]
     |> Proof.refine (Method.primitive_text (K (Goal.init (cert thesis)))) |> Seq.hd
   end;
 
--- a/src/Pure/Isar/outer_parse.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/outer_parse.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -61,12 +61,12 @@
   val list: (token list -> 'a * token list) -> token list -> 'a list * token list
   val list1: (token list -> 'a * token list) -> token list -> 'a list * token list
   val name: token list -> bstring * token list
-  val binding: token list -> Name.binding * token list
+  val binding: token list -> Binding.T * token list
   val xname: token list -> xstring * token list
   val text: token list -> string * token list
   val path: token list -> Path.T * token list
   val parname: token list -> string * token list
-  val parbinding: token list -> Name.binding * token list
+  val parbinding: token list -> Binding.T * token list
   val sort: token list -> string * token list
   val arity: token list -> (string * string list * string) * token list
   val multi_arity: token list -> (string list * string list * string) * token list
@@ -81,11 +81,11 @@
   val opt_mixfix': token list -> mixfix * token list
   val where_: token list -> string * token list
   val const: token list -> (string * string * mixfix) * token list
-  val params: token list -> (Name.binding * string option) list * token list
-  val simple_fixes: token list -> (Name.binding * string option) list * token list
-  val fixes: token list -> (Name.binding * string option * mixfix) list * token list
-  val for_fixes: token list -> (Name.binding * string option * mixfix) list * token list
-  val for_simple_fixes: token list -> (Name.binding * string option) list * token list
+  val params: token list -> (Binding.T * string option) list * token list
+  val simple_fixes: token list -> (Binding.T * string option) list * token list
+  val fixes: token list -> (Binding.T * string option * mixfix) list * token list
+  val for_fixes: token list -> (Binding.T * string option * mixfix) list * token list
+  val for_simple_fixes: token list -> (Binding.T * string option) list * token list
   val ML_source: token list -> (SymbolPos.text * Position.T) * token list
   val doc_source: token list -> (SymbolPos.text * Position.T) * token list
   val properties: token list -> Properties.T * token list
@@ -228,13 +228,13 @@
 (* names and text *)
 
 val name = group "name declaration" (short_ident || sym_ident || string || number);
-val binding = position name >> Binding.binding_pos;
+val binding = position name >> Binding.name_pos;
 val xname = group "name reference" (short_ident || long_ident || sym_ident || string || number);
 val text = group "text" (short_ident || long_ident || sym_ident || string || number || verbatim);
 val path = group "file name/path specification" name >> Path.explode;
 
 val parname = Scan.optional ($$$ "(" |-- name --| $$$ ")") "";
-val parbinding = Scan.optional ($$$ "(" |-- binding --| $$$ ")") Name.no_binding;
+val parbinding = Scan.optional ($$$ "(" |-- binding --| $$$ ")") Binding.empty;
 
 
 (* sorts and arities *)
--- a/src/Pure/Isar/proof.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/proof.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -45,27 +45,27 @@
   val match_bind_i: (term list * term) list -> state -> state
   val let_bind: (string list * string) list -> state -> state
   val let_bind_i: (term list * term) list -> state -> state
-  val fix: (Name.binding * string option * mixfix) list -> state -> state
-  val fix_i: (Name.binding * typ option * mixfix) list -> state -> state
+  val fix: (Binding.T * string option * mixfix) list -> state -> state
+  val fix_i: (Binding.T * typ option * mixfix) list -> state -> state
   val assm: Assumption.export ->
     (Attrib.binding * (string * string list) list) list -> state -> state
   val assm_i: Assumption.export ->
-    ((Name.binding * attribute list) * (term * term list) list) list -> state -> state
+    ((Binding.T * attribute list) * (term * term list) list) list -> state -> state
   val assume: (Attrib.binding * (string * string list) list) list -> state -> state
-  val assume_i: ((Name.binding * attribute list) * (term * term list) list) list ->
+  val assume_i: ((Binding.T * attribute list) * (term * term list) list) list ->
     state -> state
   val presume: (Attrib.binding * (string * string list) list) list -> state -> state
-  val presume_i: ((Name.binding * attribute list) * (term * term list) list) list ->
+  val presume_i: ((Binding.T * attribute list) * (term * term list) list) list ->
     state -> state
-  val def: (Attrib.binding * ((Name.binding * mixfix) * (string * string list))) list ->
+  val def: (Attrib.binding * ((Binding.T * mixfix) * (string * string list))) list ->
     state -> state
-  val def_i: ((Name.binding * attribute list) *
-    ((Name.binding * mixfix) * (term * term list))) list -> state -> state
+  val def_i: ((Binding.T * attribute list) *
+    ((Binding.T * mixfix) * (term * term list))) list -> state -> state
   val chain: state -> state
   val chain_facts: thm list -> state -> state
   val get_thmss: state -> (Facts.ref * Attrib.src list) list -> thm list
   val note_thmss: (Attrib.binding * (Facts.ref * Attrib.src list) list) list -> state -> state
-  val note_thmss_i: ((Name.binding * attribute list) *
+  val note_thmss_i: ((Binding.T * attribute list) *
     (thm list * attribute list) list) list -> state -> state
   val from_thmss: ((Facts.ref * Attrib.src list) list) list -> state -> state
   val from_thmss_i: ((thm list * attribute list) list) list -> state -> state
@@ -89,7 +89,7 @@
     (theory -> 'a -> attribute) ->
     (context * 'b list -> context * (term list list * (context -> context))) ->
     string -> Method.text option -> (thm list list -> state -> state Seq.seq) ->
-    ((Name.binding * 'a list) * 'b) list -> state -> state
+    ((Binding.T * 'a list) * 'b) list -> state -> state
   val local_qed: Method.text option * bool -> state -> state Seq.seq
   val theorem: Method.text option -> (thm list list -> context -> context) ->
     (string * string list) list list -> context -> state
@@ -109,13 +109,13 @@
   val have: Method.text option -> (thm list list -> state -> state Seq.seq) ->
     (Attrib.binding * (string * string list) list) list -> bool -> state -> state
   val have_i: Method.text option -> (thm list list -> state -> state Seq.seq) ->
-    ((Name.binding * attribute list) * (term * term list) list) list -> bool -> state -> state
+    ((Binding.T * attribute list) * (term * term list) list) list -> bool -> state -> state
   val show: Method.text option -> (thm list list -> state -> state Seq.seq) ->
     (Attrib.binding * (string * string list) list) list -> bool -> state -> state
   val show_i: Method.text option -> (thm list list -> state -> state Seq.seq) ->
-    ((Name.binding * attribute list) * (term * term list) list) list -> bool -> state -> state
+    ((Binding.T * attribute list) * (term * term list) list) list -> bool -> state -> state
   val is_relevant: state -> bool
-  val future_proof: (state -> context) -> state -> context
+  val future_proof: (state -> ('a * context) future) -> state -> 'a future * context
 end;
 
 structure Proof: PROOF =
@@ -617,7 +617,7 @@
 
 (* note etc. *)
 
-fun no_binding args = map (pair (Name.no_binding, [])) args;
+fun no_binding args = map (pair (Binding.empty, [])) args;
 
 local
 
@@ -645,7 +645,7 @@
 val local_results =
   gen_thmss (ProofContext.note_thmss_i "") (K []) I I (K I) o map (apsnd Thm.simple_fact);
 
-fun get_thmss state srcs = the_facts (note_thmss [((Name.no_binding, []), srcs)] state);
+fun get_thmss state srcs = the_facts (note_thmss [((Binding.empty, []), srcs)] state);
 
 end;
 
@@ -689,14 +689,14 @@
     val atts = map (prep_att (theory_of state)) raw_atts;
     val (asms, state') = state |> map_context_result (fn ctxt =>
       ctxt |> ProofContext.apply_case (ProofContext.get_case ctxt name xs));
-    val assumptions = asms |> map (fn (a, ts) => ((Name.binding a, atts), map (rpair []) ts));
+    val assumptions = asms |> map (fn (a, ts) => ((Binding.name a, atts), map (rpair []) ts));
   in
     state'
     |> map_context (ProofContext.qualified_names #> ProofContext.no_base_names)
     |> assume_i assumptions
     |> add_binds_i AutoBind.no_facts
     |> map_context (ProofContext.restore_naming (context_of state))
-    |> `the_facts |-> (fn thms => note_thmss_i [((Name.binding name, []), [(thms, [])])])
+    |> `the_facts |-> (fn thms => note_thmss_i [((Binding.name name, []), [(thms, [])])])
   end;
 
 in
@@ -990,7 +990,7 @@
   not (is_initial state) orelse
   schematic_goal state;
 
-fun future_proof make_proof state =
+fun future_proof proof state =
   let
     val _ = is_relevant state andalso error "Cannot fork relevant proof";
 
@@ -1002,18 +1002,21 @@
     val statement' = (kind, [[], [Thm.term_of cprop_protected]], cprop_protected);
     val goal' = Thm.adjust_maxidx_thm ~1 (Drule.protectI RS Goal.init cprop_protected);
     fun after_qed' [[th]] = ProofContext.put_thms false (AutoBind.thisN, SOME [Goal.conclude th]);
-    val this_name = ProofContext.full_name (ProofContext.reset_naming goal_ctxt) AutoBind.thisN;
+    val this_name = ProofContext.full_bname (ProofContext.reset_naming goal_ctxt) AutoBind.thisN;
 
-    fun make_result () = state
+    val result_ctxt =
+      state
       |> map_contexts (Variable.auto_fixes prop)
       |> map_goal I (K (statement', messages, using, goal', before_qed, (#1 after_qed, after_qed')))
-      |> make_proof
-      |> (fn result_ctxt => ProofContext.get_fact_single result_ctxt (Facts.named this_name));
-    val finished_goal = Goal.protect (Goal.future_result goal_ctxt make_result prop);
-  in
-    state
-    |> map_goal I (K (statement, messages, using, finished_goal, NONE, after_qed))
-    |> global_done_proof
-  end;
+      |> proof;
+
+    val thm = result_ctxt |> Future.map (fn (_, ctxt) =>
+      ProofContext.get_fact_single ctxt (Facts.named this_name));
+    val finished_goal = Goal.protect (Goal.future_result goal_ctxt thm prop);
+    val state' =
+      state
+      |> map_goal I (K (statement, messages, using, finished_goal, NONE, after_qed))
+      |> global_done_proof;
+  in (Future.map #1 result_ctxt, state') end;
 
 end;
--- a/src/Pure/Isar/proof_context.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/proof_context.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -23,8 +23,8 @@
   val abbrev_mode: Proof.context -> bool
   val set_stmt: bool -> Proof.context -> Proof.context
   val naming_of: Proof.context -> NameSpace.naming
-  val full_name: Proof.context -> bstring -> string
-  val full_binding: Proof.context -> Name.binding -> string
+  val full_name: Proof.context -> Binding.T -> string
+  val full_bname: Proof.context -> bstring -> string
   val consts_of: Proof.context -> Consts.T
   val const_syntax_name: Proof.context -> string -> string
   val the_const_constraint: Proof.context -> string -> typ
@@ -105,27 +105,27 @@
   val restore_naming: Proof.context -> Proof.context -> Proof.context
   val reset_naming: Proof.context -> Proof.context
   val note_thmss: string ->
-    ((Name.binding * attribute list) * (Facts.ref * attribute list) list) list ->
+    ((Binding.T * attribute list) * (Facts.ref * attribute list) list) list ->
       Proof.context -> (string * thm list) list * Proof.context
   val note_thmss_i: string ->
-    ((Name.binding * attribute list) * (thm list * attribute list) list) list ->
+    ((Binding.T * attribute list) * (thm list * attribute list) list) list ->
       Proof.context -> (string * thm list) list * Proof.context
   val put_thms: bool -> string * thm list option -> Proof.context -> Proof.context
-  val read_vars: (Name.binding * string option * mixfix) list -> Proof.context ->
-    (Name.binding * typ option * mixfix) list * Proof.context
-  val cert_vars: (Name.binding * typ option * mixfix) list -> Proof.context ->
-    (Name.binding * typ option * mixfix) list * Proof.context
-  val add_fixes: (Name.binding * string option * mixfix) list ->
+  val read_vars: (Binding.T * string option * mixfix) list -> Proof.context ->
+    (Binding.T * typ option * mixfix) list * Proof.context
+  val cert_vars: (Binding.T * typ option * mixfix) list -> Proof.context ->
+    (Binding.T * typ option * mixfix) list * Proof.context
+  val add_fixes: (Binding.T * string option * mixfix) list ->
     Proof.context -> string list * Proof.context
-  val add_fixes_i: (Name.binding * typ option * mixfix) list ->
+  val add_fixes_i: (Binding.T * typ option * mixfix) list ->
     Proof.context -> string list * Proof.context
   val auto_fixes: Proof.context * (term list list * 'a) -> Proof.context * (term list list * 'a)
   val bind_fixes: string list -> Proof.context -> (term -> term) * Proof.context
   val add_assms: Assumption.export ->
-    ((Name.binding * attribute list) * (string * string list) list) list ->
+    ((Binding.T * attribute list) * (string * string list) list) list ->
     Proof.context -> (string * thm list) list * Proof.context
   val add_assms_i: Assumption.export ->
-    ((Name.binding * attribute list) * (term * term list) list) list ->
+    ((Binding.T * attribute list) * (term * term list) list) list ->
     Proof.context -> (string * thm list) list * Proof.context
   val add_cases: bool -> (string * RuleCases.T option) list -> Proof.context -> Proof.context
   val apply_case: RuleCases.T -> Proof.context -> (string * term list) list * Proof.context
@@ -135,7 +135,7 @@
     Context.generic -> Context.generic
   val add_const_constraint: string * typ option -> Proof.context -> Proof.context
   val add_abbrev: string -> Properties.T ->
-    Name.binding * term -> Proof.context -> (term * term) * Proof.context
+    Binding.T * term -> Proof.context -> (term * term) * Proof.context
   val revert_abbrev: string -> string -> Proof.context -> Proof.context
   val verbose: bool ref
   val setmp_verbose: ('a -> 'b) -> 'a -> 'b
@@ -247,8 +247,8 @@
 
 val naming_of = #naming o rep_context;
 
-val full_name = NameSpace.full o naming_of;
-val full_binding = NameSpace.full_binding o naming_of;
+val full_name = NameSpace.full_name o naming_of;
+fun full_bname thy = NameSpace.full_name (naming_of thy) o Binding.name;
 
 val syntax_of = #syntax o rep_context;
 val syn_of = LocalSyntax.syn_of o syntax_of;
@@ -965,14 +965,14 @@
 
 local
 
-fun update_thms _ (b, NONE) ctxt = ctxt |> map_facts (Facts.del (full_binding ctxt b))
+fun update_thms _ (b, NONE) ctxt = ctxt |> map_facts (Facts.del (full_name ctxt b))
   | update_thms do_props (b, SOME ths) ctxt = ctxt |> map_facts
       (Facts.add_local do_props (naming_of ctxt) (b, ths) #> snd);
 
 fun gen_note_thmss get k = fold_map (fn ((b, more_attrs), raw_facts) => fn ctxt =>
   let
     val pos = Binding.pos_of b;
-    val name = full_binding ctxt b;
+    val name = full_name ctxt b;
     val _ = ContextPosition.report_visible ctxt (Markup.local_fact_decl name) pos;
 
     val facts = PureThy.name_thmss false pos name (map (apfst (get ctxt)) raw_facts);
@@ -991,7 +991,7 @@
 fun put_thms do_props thms ctxt = ctxt
   |> map_naming (K local_naming)
   |> ContextPosition.set_visible false
-  |> update_thms do_props (apfst Name.binding thms)
+  |> update_thms do_props (apfst Binding.name thms)
   |> ContextPosition.restore_visible ctxt
   |> restore_naming ctxt;
 
@@ -1012,7 +1012,7 @@
 fun prep_vars prep_typ internal =
   fold_map (fn (raw_b, raw_T, raw_mx) => fn ctxt =>
     let
-      val raw_x = Name.name_of raw_b;
+      val raw_x = Binding.base_name raw_b;
       val (x, mx) = Syntax.const_mixfix raw_x raw_mx;
       val _ = Syntax.is_identifier (no_skolem internal x) orelse
         error ("Illegal variable name: " ^ quote x);
@@ -1021,7 +1021,7 @@
         if internal then T
         else Type.no_tvars T handle TYPE (msg, _, _) => error msg;
       val opt_T = Option.map (cond_tvars o cert_typ ctxt o prep_typ ctxt) raw_T;
-      val var = (Name.map_name (K x) raw_b, opt_T, mx);
+      val var = (Binding.map_base (K x) raw_b, opt_T, mx);
     in (var, ctxt |> declare_var (x, opt_T, mx) |> #2) end);
 
 in
@@ -1095,7 +1095,7 @@
 fun add_abbrev mode tags (b, raw_t) ctxt =
   let
     val t0 = cert_term (ctxt |> set_mode mode_abbrev) raw_t
-      handle ERROR msg => cat_error msg ("in constant abbreviation " ^ quote (Name.display b));
+      handle ERROR msg => cat_error msg ("in constant abbreviation " ^ quote (Binding.display b));
     val [t] = Variable.exportT_terms (Variable.declare_term t0 ctxt) ctxt [t0];
     val ((lhs, rhs), consts') = consts_of ctxt
       |> Consts.abbreviate (Syntax.pp ctxt) (tsig_of ctxt) (naming_of ctxt) mode tags (b, t);
@@ -1122,7 +1122,7 @@
 fun gen_fixes prep raw_vars ctxt =
   let
     val (vars, _) = prep raw_vars ctxt;
-    val (xs', ctxt') = Variable.add_fixes (map (Name.name_of o #1) vars) ctxt;
+    val (xs', ctxt') = Variable.add_fixes (map (Binding.base_name o #1) vars) ctxt;
     val ctxt'' =
       ctxt'
       |> fold_map declare_var (map2 (fn x' => fn (_, T, mx) => (x', T, mx)) xs' vars)
@@ -1146,7 +1146,7 @@
 
 fun bind_fixes xs ctxt =
   let
-    val (_, ctxt') = ctxt |> add_fixes_i (map (fn x => (Name.binding x, NONE, NoSyn)) xs);
+    val (_, ctxt') = ctxt |> add_fixes_i (map (fn x => (Binding.name x, NONE, NoSyn)) xs);
     fun bind (t as Free (x, T)) =
           if member (op =) xs x then
             (case lookup_skolem ctxt' x of SOME x' => Free (x', T) | NONE => t)
--- a/src/Pure/Isar/rule_insts.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/rule_insts.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -284,7 +284,7 @@
         val (param_names, ctxt') = ctxt
           |> Variable.declare_thm thm
           |> Thm.fold_terms Variable.declare_constraints st
-          |> ProofContext.add_fixes_i (map (fn (x, T) => (Name.binding x, SOME T, NoSyn)) params);
+          |> ProofContext.add_fixes_i (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) params);
 
         (* Process type insts: Tinsts_env *)
         fun absent xi = error
--- a/src/Pure/Isar/spec_parse.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/spec_parse.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -15,14 +15,14 @@
   val opt_thm_name: string -> token list -> Attrib.binding * token list
   val spec: token list -> (Attrib.binding * string list) * token list
   val named_spec: token list -> (Attrib.binding * string list) * token list
-  val spec_name: token list -> ((Name.binding * string) * Attrib.src list) * token list
-  val spec_opt_name: token list -> ((Name.binding * string) * Attrib.src list) * token list
+  val spec_name: token list -> ((Binding.T * string) * Attrib.src list) * token list
+  val spec_opt_name: token list -> ((Binding.T * string) * Attrib.src list) * token list
   val xthm: token list -> (Facts.ref * Attrib.src list) * token list
   val xthms1: token list -> (Facts.ref * Attrib.src list) list * token list
   val name_facts: token list ->
     (Attrib.binding * (Facts.ref * Attrib.src list) list) list * token list
   val locale_mixfix: token list -> mixfix * token list
-  val locale_fixes: token list -> (Name.binding * string option * mixfix) list * token list
+  val locale_fixes: token list -> (Binding.T * string option * mixfix) list * token list
   val locale_insts: token list -> (string option list * (Attrib.binding * string) list) * token list
   val class_expr: token list -> string list * token list
   val locale_expr: token list -> Locale.expr * token list
@@ -33,8 +33,8 @@
     (Element.context list * Element.statement) * OuterLex.token list
   val statement_keyword: token list -> string * token list
   val specification: token list ->
-    (Name.binding *
-      ((Attrib.binding * string list) list * (Name.binding * string option) list)) list * token list
+    (Binding.T *
+      ((Attrib.binding * string list) list * (Binding.T * string option) list)) list * token list
 end;
 
 structure SpecParse: SPEC_PARSE =
@@ -53,8 +53,8 @@
 fun thm_name s = P.binding -- opt_attribs --| P.$$$ s;
 
 fun opt_thm_name s =
-  Scan.optional ((P.binding -- opt_attribs || attribs >> pair Name.no_binding) --| P.$$$ s)
-    Attrib.no_binding;
+  Scan.optional ((P.binding -- opt_attribs || attribs >> pair Binding.empty) --| P.$$$ s)
+    Attrib.empty_binding;
 
 val spec = opt_thm_name ":" -- Scan.repeat1 P.prop;
 val named_spec = thm_name ":" -- Scan.repeat1 P.prop;
--- a/src/Pure/Isar/specification.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/specification.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -9,33 +9,33 @@
 signature SPECIFICATION =
 sig
   val print_consts: local_theory -> (string * typ -> bool) -> (string * typ) list -> unit
-  val check_specification: (Name.binding * typ option * mixfix) list ->
+  val check_specification: (Binding.T * typ option * mixfix) list ->
     (Attrib.binding * term list) list list -> Proof.context ->
-    (((Name.binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
-  val read_specification: (Name.binding * string option * mixfix) list ->
+    (((Binding.T * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
+  val read_specification: (Binding.T * string option * mixfix) list ->
     (Attrib.binding * string list) list list -> Proof.context ->
-    (((Name.binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
-  val check_free_specification: (Name.binding * typ option * mixfix) list ->
+    (((Binding.T * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
+  val check_free_specification: (Binding.T * typ option * mixfix) list ->
     (Attrib.binding * term list) list -> Proof.context ->
-    (((Name.binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
-  val read_free_specification: (Name.binding * string option * mixfix) list ->
+    (((Binding.T * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
+  val read_free_specification: (Binding.T * string option * mixfix) list ->
     (Attrib.binding * string list) list -> Proof.context ->
-    (((Name.binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
-  val axiomatization: (Name.binding * typ option * mixfix) list ->
+    (((Binding.T * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
+  val axiomatization: (Binding.T * typ option * mixfix) list ->
     (Attrib.binding * term list) list -> theory ->
     (term list * (string * thm list) list) * theory
-  val axiomatization_cmd: (Name.binding * string option * mixfix) list ->
+  val axiomatization_cmd: (Binding.T * string option * mixfix) list ->
     (Attrib.binding * string list) list -> theory ->
     (term list * (string * thm list) list) * theory
   val definition:
-    (Name.binding * typ option * mixfix) option * (Attrib.binding * term) ->
+    (Binding.T * typ option * mixfix) option * (Attrib.binding * term) ->
     local_theory -> (term * (string * thm)) * local_theory
   val definition_cmd:
-    (Name.binding * string option * mixfix) option * (Attrib.binding * string) ->
+    (Binding.T * string option * mixfix) option * (Attrib.binding * string) ->
     local_theory -> (term * (string * thm)) * local_theory
-  val abbreviation: Syntax.mode -> (Name.binding * typ option * mixfix) option * term ->
+  val abbreviation: Syntax.mode -> (Binding.T * typ option * mixfix) option * term ->
     local_theory -> local_theory
-  val abbreviation_cmd: Syntax.mode -> (Name.binding * string option * mixfix) option * string ->
+  val abbreviation_cmd: Syntax.mode -> (Binding.T * string option * mixfix) option * string ->
     local_theory -> local_theory
   val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory
   val notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory
@@ -153,7 +153,7 @@
 fun gen_axioms do_print prep raw_vars raw_specs thy =
   let
     val ((vars, specs), _) = prep raw_vars [raw_specs] (ProofContext.init thy);
-    val xs = map (fn ((b, T), _) => (Name.name_of b, T)) vars;
+    val xs = map (fn ((b, T), _) => (Binding.base_name b, T)) vars;
 
     (*consts*)
     val (consts, consts_thy) = thy |> fold_map (Theory.specify_const []) vars;
@@ -161,7 +161,7 @@
 
     (*axioms*)
     val (axioms, axioms_thy) = consts_thy |> fold_map (fn ((b, atts), props) =>
-        fold_map Thm.add_axiom (PureThy.name_multi (Name.name_of b) (map subst props))
+        fold_map Thm.add_axiom (PureThy.name_multi (Binding.base_name b) (map subst props))
         #>> (fn ths => ((b, atts), [(map Drule.standard' ths, [])]))) specs;
     val (facts, thy') = axioms_thy |> PureThy.note_thmss Thm.axiomK
       (Attrib.map_facts (Attrib.attribute_i axioms_thy) axioms);
@@ -181,19 +181,19 @@
     val (vars, [((raw_name, atts), [prop])]) =
       fst (prep (the_list raw_var) [(raw_a, [raw_prop])] lthy);
     val (((x, T), rhs), prove) = LocalDefs.derived_def lthy true prop;
-    val name = Name.map_name (Thm.def_name_optional x) raw_name;
+    val name = Binding.map_base (Thm.def_name_optional x) raw_name;
     val var =
       (case vars of
-        [] => (Name.binding x, NoSyn)
+        [] => (Binding.name x, NoSyn)
       | [((b, _), mx)] =>
           let
-            val y = Name.name_of b;
+            val y = Binding.base_name b;
             val _ = x = y orelse
               error ("Head of definition " ^ quote x ^ " differs from declaration " ^ quote y ^
                 Position.str_of (Binding.pos_of b));
           in (b, mx) end);
     val ((lhs, (_, th)), lthy2) = lthy |> LocalTheory.define Thm.definitionK
-        (var, ((Name.map_name (suffix "_raw") name, []), rhs));
+        (var, ((Binding.map_base (suffix "_raw") name, []), rhs));
     val ((def_name, [th']), lthy3) = lthy2 |> LocalTheory.note Thm.definitionK
         ((name, Code.add_default_eqn_attrib :: atts), [prove lthy2 th]);
 
@@ -217,10 +217,10 @@
     val ((x, T), rhs) = LocalDefs.abs_def (#2 (LocalDefs.cert_def lthy prop));
     val var =
       (case vars of
-        [] => (Name.binding x, NoSyn)
+        [] => (Binding.name x, NoSyn)
       | [((b, _), mx)] =>
           let
-            val y = Name.name_of b;
+            val y = Binding.base_name b;
             val _ = x = y orelse
               error ("Head of abbreviation " ^ quote x ^ " differs from declaration " ^ quote y ^
                 Position.str_of (Binding.pos_of b));
@@ -281,11 +281,11 @@
   | Element.Obtains obtains =>
       let
         val case_names = obtains |> map_index (fn (i, (b, _)) =>
-          let val name = Name.name_of b
+          let val name = Binding.base_name b
           in if name = "" then string_of_int (i + 1) else name end);
         val constraints = obtains |> map (fn (_, (vars, _)) =>
           Element.Constrains
-            (vars |> map_filter (fn (x, SOME T) => SOME (Name.name_of x, T) | _ => NONE)));
+            (vars |> map_filter (fn (x, SOME T) => SOME (Binding.base_name x, T) | _ => NONE)));
 
         val raw_propp = obtains |> map (fn (_, (_, props)) => map (rpair []) props);
         val (propp, elems_ctxt) = prep_stmt (elems @ constraints) raw_propp ctxt;
@@ -295,7 +295,7 @@
         fun assume_case ((name, (vars, _)), asms) ctxt' =
           let
             val bs = map fst vars;
-            val xs = map Name.name_of bs;
+            val xs = map Binding.base_name bs;
             val props = map fst asms;
             val (Ts, _) = ctxt'
               |> fold Variable.declare_term props
@@ -312,13 +312,13 @@
         val atts = map (Attrib.internal o K)
           [RuleCases.consumes (~ (length obtains)), RuleCases.case_names case_names];
         val prems = subtract_prems ctxt elems_ctxt;
-        val stmt = [((Name.no_binding, atts), [(thesis, [])])];
+        val stmt = [((Binding.empty, atts), [(thesis, [])])];
 
         val (facts, goal_ctxt) = elems_ctxt
-          |> (snd o ProofContext.add_fixes_i [(Name.binding AutoBind.thesisN, NONE, NoSyn)])
+          |> (snd o ProofContext.add_fixes_i [(Binding.name AutoBind.thesisN, NONE, NoSyn)])
           |> fold_map assume_case (obtains ~~ propp)
           |-> (fn ths => ProofContext.note_thmss_i Thm.assumptionK
-                [((Name.binding Obtain.thatN, []), [(ths, [])])] #> #2 #> pair ths);
+                [((Binding.name Obtain.thatN, []), [(ths, [])])] #> #2 #> pair ths);
       in ((prems, stmt, facts), goal_ctxt) end);
 
 structure TheoremHook = GenericDataFun
@@ -348,7 +348,7 @@
         lthy
         |> LocalTheory.notes kind (map2 (fn (a, _) => fn ths => (a, [(ths, [])])) stmt results')
         |> (fn (res, lthy') =>
-          if Binding.is_nothing name andalso null atts then
+          if Binding.is_empty name andalso null atts then
             (ProofDisplay.print_results true lthy' ((kind, ""), res); lthy')
           else
             let
@@ -361,7 +361,7 @@
   in
     goal_ctxt
     |> ProofContext.note_thmss_i Thm.assumptionK
-      [((Name.binding AutoBind.assmsN, []), [(prems, [])])]
+      [((Binding.name AutoBind.assmsN, []), [(prems, [])])]
     |> snd
     |> Proof.theorem_i before_qed after_qed' (map snd stmt)
     |> Proof.refine_insert facts
--- a/src/Pure/Isar/theory_target.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/theory_target.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -59,9 +59,9 @@
   let
     val thy = ProofContext.theory_of ctxt;
     val target_name = (if is_class then "class " else "locale ") ^ locale_extern thy target;
-    val fixes = map (fn (x, T) => (Name.binding x, SOME T, NoSyn))
+    val fixes = map (fn (x, T) => (Binding.name x, SOME T, NoSyn))
       (#1 (ProofContext.inferred_fixes ctxt));
-    val assumes = map (fn A => (Attrib.no_binding, [(Thm.term_of A, [])]))
+    val assumes = map (fn A => (Attrib.empty_binding, [(Thm.term_of A, [])]))
       (Assumption.assms_of ctxt);
     val elems =
       (if null fixes then [] else [Element.Fixes fixes]) @
@@ -163,11 +163,9 @@
 fun notes (Target {target, is_locale, is_class, ...}) kind facts lthy =
   let
     val thy = ProofContext.theory_of lthy;
-    val full = LocalTheory.full_name lthy;
-
     val facts' = facts
-      |> map (fn (a, bs) =>
-        (a, PureThy.burrow_fact (PureThy.name_multi (full (Name.name_of (fst a)))) bs))
+      |> map (fn (a, bs) => (a, PureThy.burrow_fact (PureThy.name_multi
+          (LocalTheory.full_name lthy (fst a))) bs))
       |> PureThy.map_facts (import_export_proof lthy);
     val local_facts = PureThy.map_facts #1 facts'
       |> Attrib.map_facts (Attrib.attribute_i thy);
@@ -195,14 +193,14 @@
 
 fun locale_const (Target {target, is_class, ...}) (prmode as (mode, _)) tags ((b, mx), rhs) phi =
   let
-    val b' = Morphism.name phi b;
+    val b' = Morphism.binding phi b;
     val rhs' = Morphism.term phi rhs;
     val legacy_arg = (b', Term.close_schematic_term (Logic.legacy_varify rhs'));
     val arg = (b', Term.close_schematic_term rhs');
     val similar_body = Type.similar_types (rhs, rhs');
     (* FIXME workaround based on educated guess *)
-    val (prefix', _) = Binding.dest_binding b';
-    val class_global = Name.name_of b = Name.name_of b'
+    val (prefix', _) = Binding.dest b';
+    val class_global = Binding.base_name b = Binding.base_name b'
       andalso not (null prefix')
       andalso (fst o snd o split_last) prefix' = Class.class_prefix target;
   in
@@ -221,7 +219,7 @@
 
 fun declare_const (ta as Target {target, is_locale, is_class, ...}) depends ((b, T), mx) lthy =
   let
-    val c = Name.name_of b;
+    val c = Binding.base_name b;
     val tags = LocalTheory.group_position_of lthy;
     val xs = filter depends (#1 (ProofContext.inferred_fixes (LocalTheory.target_of lthy)));
     val U = map #2 xs ---> T;
@@ -254,7 +252,7 @@
 
 fun abbrev (ta as Target {target, is_locale, is_class, ...}) prmode ((b, mx), t) lthy =
   let
-    val c = Name.name_of b;
+    val c = Binding.base_name b;
     val tags = LocalTheory.group_position_of lthy;
     val thy_ctxt = ProofContext.init (ProofContext.theory_of lthy);
     val target_ctxt = LocalTheory.target_of lthy;
@@ -291,8 +289,8 @@
     val thy = ProofContext.theory_of lthy;
     val thy_ctxt = ProofContext.init thy;
 
-    val c = Name.name_of b;
-    val name' = Name.map_name (Thm.def_name_optional c) name;
+    val c = Binding.base_name b;
+    val name' = Binding.map_base (Thm.def_name_optional c) name;
     val (rhs', rhs_conv) =
       LocalDefs.export_cterm lthy thy_ctxt (Thm.cterm_of thy rhs) |>> Thm.term_of;
     val xs = Variable.add_fixed (LocalTheory.target_of lthy) rhs' [];
@@ -312,7 +310,7 @@
           then (fn name => fn eq => Thm.add_def false false (name, Logic.mk_equals eq))
           else (fn name => fn (Const (c, _), rhs) => AxClass.define_overloaded name (c, rhs)));
     val (global_def, lthy3) = lthy2
-      |> LocalTheory.theory_result (define_const (Name.name_of name') (lhs', rhs'));
+      |> LocalTheory.theory_result (define_const (Binding.base_name name') (lhs', rhs'));
     val def = LocalDefs.trans_terms lthy3
       [(*c == global.c xs*)     local_def,
        (*global.c xs == rhs'*)  global_def,
--- a/src/Pure/Isar/toplevel.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Isar/toplevel.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -691,45 +691,62 @@
 
 (* excursion of units, consisting of commands with proof *)
 
+structure States = ProofDataFun
+(
+  type T = state list future option;
+  fun init _ = NONE;
+);
+
 fun command_result tr st =
   let val st' = command tr st
   in (st', st') end;
 
-fun unit_result immediate (tr, proof_trs) st =
+fun proof_result immediate (tr, proof_trs) st =
   let val st' = command tr st in
     if immediate orelse null proof_trs orelse
       not (can proof_of st') orelse Proof.is_relevant (proof_of st')
     then
       let val (states, st'') = fold_map command_result proof_trs st'
-      in (fn () => (tr, st') :: (proof_trs ~~ states), st'') end
+      in (Lazy.value ((tr, st') :: (proof_trs ~~ states)), st'') end
     else
       let
         val (body_trs, end_tr) = split_last proof_trs;
-        val body_states = ref ([]: state list);
         val finish = Context.Theory o ProofContext.theory_of;
-        fun make_result prf =
-          let val (states, State (result_node, _)) =
-            (case st' of State (SOME (Proof (_, (_, orig_gthy)), exit), prev)
-              => State (SOME (Proof (ProofNode.init prf, (finish, orig_gthy)), exit), prev))
-            |> fold_map command_result body_trs
-            ||> command (end_tr |> set_print false)
-          in body_states := states; presentation_context (Option.map #1 result_node) NONE end;
-        val st'' = st'
-          |> command (end_tr |> reset_trans |> end_proof (K (Proof.future_proof make_result)));
-      in (fn () => (tr, st') :: (body_trs ~~ ! body_states) @ [(end_tr, st'')], st'') end
+
+        val future_proof = Proof.future_proof
+          (fn prf =>
+            Future.fork_background (fn () =>
+              let val (states, State (result_node, _)) =
+                (case st' of State (SOME (Proof (_, (_, orig_gthy)), exit), prev)
+                  => State (SOME (Proof (ProofNode.init prf, (finish, orig_gthy)), exit), prev))
+                |> fold_map command_result body_trs
+                ||> command (end_tr |> set_print false);
+              in (states, presentation_context (Option.map #1 result_node) NONE) end))
+          #> (fn (states, ctxt) => States.put (SOME states) ctxt);
+
+        val st'' = st' |> command (end_tr |> reset_trans |> end_proof (K future_proof));
+
+        val states =
+          (case States.get (presentation_context (SOME (node_of st'')) NONE) of
+            NONE => sys_error ("No future states for " ^ name_of tr ^ Position.str_of (pos_of tr))
+          | SOME states => states);
+        val result = Lazy.lazy
+          (fn () => (tr, st') :: (body_trs ~~ Future.join states) @ [(end_tr, st'')]);
+
+      in (result, st'') end
   end;
 
-fun excursion input =
+fun excursion input = exception_trace (fn () =>
   let
     val end_pos = if null input then error "No input" else pos_of (fst (List.last input));
 
     val immediate = not (Future.enabled ());
-    val (future_results, end_state) = fold_map (unit_result immediate) input toplevel;
+    val (future_results, end_state) = fold_map (proof_result immediate) input toplevel;
     val _ =
       (case end_state of
-        State (NONE, SOME (Theory (Context.Theory thy, _), _)) => Thm.join_futures thy
+        State (NONE, SOME (Theory (Context.Theory _, _), _)) => ()
       | _ => error "Unfinished development at end of input");
-    val results = maps (fn res => res ()) future_results;
-  in (results, fn () => ignore (command (commit_exit end_pos) end_state)) end;
+    val results = maps Lazy.force future_results;
+  in (results, fn () => ignore (command (commit_exit end_pos) end_state)) end);
 
 end;
--- a/src/Pure/ML-Systems/install_pp_polyml.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/ML-Systems/install_pp_polyml.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -4,13 +4,13 @@
 Extra toplevel pretty-printing for Poly/ML.
 *)
 
-install_pp (fn (str, _, _, _) => fn depth => fn (print: 'a * int -> unit) => fn (x: 'a Future.T) =>
+install_pp (fn (str, _, _, _) => fn depth => fn (print: 'a * int -> unit) => fn (x: 'a future) =>
   (case Future.peek x of
     NONE => str "<future>"
   | SOME (Exn.Exn _) => str "<failed>"
   | SOME (Exn.Result y) => print (y, depth)));
 
-install_pp (fn (str, _, _, _) => fn depth => fn (print: 'a * int -> unit) => fn (x: 'a Lazy.T) =>
+install_pp (fn (str, _, _, _) => fn depth => fn (print: 'a * int -> unit) => fn (x: 'a lazy) =>
   (case Lazy.peek x of
     NONE => str "<lazy>"
   | SOME (Exn.Exn _) => str "<failed>"
--- a/src/Pure/ROOT.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/ROOT.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ROOT.ML
-    ID:         $Id$
 
 Pure Isabelle.
 *)
--- a/src/Pure/Syntax/syntax.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Syntax/syntax.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -173,7 +173,7 @@
 
 fun remove_trtab trfuns = fold (Symtab.remove SynExt.eq_trfun) trfuns;
 
-fun update_trtab name trfuns tab = Symtab.extend (remove_trtab trfuns tab, trfuns)
+fun update_trtab name trfuns tab = fold Symtab.update_new trfuns (remove_trtab trfuns tab)
   handle Symtab.DUP c => err_dup_trfun name c;
 
 fun merge_trtabs name tab1 tab2 = Symtab.merge SynExt.eq_trfun (tab1, tab2)
--- a/src/Pure/Thy/thy_info.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Thy/thy_info.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -320,15 +320,24 @@
     val tasks = Graph.topological_order task_graph |> map_filter (fn name =>
       (case Graph.get_node task_graph name of Task body => SOME (name, body) | _ => NONE));
 
-    val group = TaskQueue.new_group ();
     fun future (name, body) tab =
       let
-        val deps = map_filter (Symtab.lookup tab) (Graph.imm_preds task_graph name);
-        val x = Future.future (SOME group) deps true body;
-      in (x, Symtab.update (name, Future.task_of x) tab) end;
-    val _ = ignore (Exn.release_all (Future.join_results (#1 (fold_map future tasks Symtab.empty))));
-    val _ = List.app (PureThy.force_proofs o get_theory o #1) tasks;
-  in () end;
+        val deps = Graph.imm_preds task_graph name
+          |> map_filter (fn parent =>
+            (case Symtab.lookup tab parent of SOME future => SOME (parent, future) | NONE => NONE));
+        fun failed (parent, future) = if can Future.join future then NONE else SOME parent;
+
+        val x = Future.fork_deps (map #2 deps) (fn () =>
+          (case map_filter failed deps of
+            [] => body ()
+          | bad => error (loader_msg
+              ("failed to load " ^ quote name ^ " (unresolved " ^ commas_quote bad ^ ")") [])));
+
+      in (x, Symtab.update (name, x) tab) end;
+
+    val thy_results = Future.join_results (#1 (fold_map future tasks Symtab.empty));
+    val proof_results = PureThy.join_proofs (map_filter (try get_theory o #1) tasks);
+  in ignore (Exn.release_all (thy_results @ proof_results)) end;
 
 local
 
@@ -589,8 +598,6 @@
 
 (* finish all theories *)
 
-fun finish () = change_thys (Graph.map_nodes
-  (fn (SOME _, SOME thy) => (NONE, (PureThy.force_proofs thy; SOME thy))
-    | (_, entry) => (NONE, entry)));
+fun finish () = change_thys (Graph.map_nodes (fn (_, entry) => (NONE, entry)));
 
 end;
--- a/src/Pure/Tools/invoke.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/Tools/invoke.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Tools/invoke.ML
-    ID:         $Id$
     Author:     Makarius
 
 Schematic invocation of locale expression in proof context.
@@ -8,9 +7,9 @@
 signature INVOKE =
 sig
   val invoke: string * Attrib.src list -> Locale.expr -> string option list ->
-    (Name.binding * string option * mixfix) list -> bool -> Proof.state -> Proof.state
+    (Binding.T * string option * mixfix) list -> bool -> Proof.state -> Proof.state
   val invoke_i: string * attribute list -> Locale.expr -> term option list ->
-    (Name.binding * typ option * mixfix) list -> bool -> Proof.state -> Proof.state
+    (Binding.T * typ option * mixfix) list -> bool -> Proof.state -> Proof.state
 end;
 
 structure Invoke: INVOKE =
@@ -60,9 +59,9 @@
         | NONE => Drule.termI)) params';
 
     val propp =
-      [((Name.no_binding, []), map (rpair [] o Logic.mk_term o Logic.mk_type) types'),
-       ((Name.no_binding, []), map (rpair [] o Logic.mk_term) params'),
-       ((Name.no_binding, []), map (rpair [] o Element.mark_witness) prems')];
+      [((Binding.empty, []), map (rpair [] o Logic.mk_term o Logic.mk_type) types'),
+       ((Binding.empty, []), map (rpair [] o Logic.mk_term) params'),
+       ((Binding.empty, []), map (rpair [] o Element.mark_witness) prems')];
     fun after_qed results =
       Proof.end_block #>
       Proof.map_context (fn ctxt =>
@@ -121,7 +120,7 @@
     (K.tag_proof K.prf_goal)
     (SpecParse.opt_thm_name ":" -- SpecParse.locale_expr -- SpecParse.locale_insts -- P.for_fixes
       >> (fn ((((name, atts), expr), (insts, _)), fixes) =>
-          Toplevel.print o Toplevel.proof' (invoke (Name.name_of name, atts) expr insts fixes)));
+          Toplevel.print o Toplevel.proof' (invoke (Binding.base_name name, atts) expr insts fixes)));
 
 end;
 
--- a/src/Pure/assumption.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/assumption.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -120,6 +120,6 @@
     val thm = export false inner outer;
     val term = export_term inner outer;
     val typ = Logic.type_map term;
-  in Morphism.morphism {name = I, var = I, typ = typ, term = term, fact = map thm} end;
+  in Morphism.morphism {binding = I, var = I, typ = typ, term = term, fact = map thm} end;
 
 end;
--- a/src/Pure/axclass.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/axclass.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -9,7 +9,7 @@
 signature AX_CLASS =
 sig
   val define_class: bstring * class list -> string list ->
-    ((Name.binding * attribute list) * term list) list -> theory -> class * theory
+    ((Binding.T * attribute list) * term list) list -> theory -> class * theory
   val add_classrel: thm -> theory -> theory
   val add_arity: thm -> theory -> theory
   val prove_classrel: class * class -> tactic -> theory -> theory
@@ -370,7 +370,7 @@
     thy
     |> Sign.sticky_prefix name_inst
     |> Sign.no_base_names
-    |> Sign.declare_const [] ((Name.binding c', T'), NoSyn)
+    |> Sign.declare_const [] ((Binding.name c', T'), NoSyn)
     |-> (fn const' as Const (c'', _) => Thm.add_def false true
           (Thm.def_name c', Logic.mk_equals (Const (c, T'), const'))
     #>> Thm.varifyT
@@ -422,7 +422,7 @@
     (* class *)
 
     val bconst = Logic.const_of_class bclass;
-    val class = Sign.full_name thy bclass;
+    val class = Sign.full_bname thy bclass;
     val super = Sign.minimize_sort thy (Sign.certify_sort thy raw_super);
 
     fun check_constraint (a, S) =
@@ -482,9 +482,9 @@
       |> Sign.add_path bconst
       |> Sign.no_base_names
       |> PureThy.note_thmss ""
-        [((Name.binding introN, []), [([Drule.standard raw_intro], [])]),
-         ((Name.binding superN, []), [(map Drule.standard raw_classrel, [])]),
-         ((Name.binding axiomsN, []),
+        [((Binding.name introN, []), [([Drule.standard raw_intro], [])]),
+         ((Binding.name superN, []), [(map Drule.standard raw_classrel, [])]),
+         ((Binding.name axiomsN, []),
            [(map (fn th => Drule.standard (class_triv RS th)) raw_axioms, [])])]
       ||> Sign.restore_naming def_thy;
 
@@ -530,7 +530,7 @@
 
 fun ax_class prep_class prep_classrel (bclass, raw_super) thy =
   let
-    val class = Sign.full_name thy bclass;
+    val class = Sign.full_bname thy bclass;
     val super = map (prep_class thy) raw_super |> Sign.minimize_sort thy;
   in
     thy
--- a/src/Pure/consts.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/consts.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -30,10 +30,10 @@
   val certify: Pretty.pp -> Type.tsig -> bool -> T -> term -> term  (*exception TYPE*)
   val typargs: T -> string * typ -> typ list
   val instance: T -> string * typ list -> typ
-  val declare: bool -> NameSpace.naming -> Properties.T -> (Name.binding * typ) -> T -> T
+  val declare: bool -> NameSpace.naming -> Properties.T -> (Binding.T * typ) -> T -> T
   val constrain: string * typ option -> T -> T
   val abbreviate: Pretty.pp -> Type.tsig -> NameSpace.naming -> string -> Properties.T ->
-    Name.binding * term -> T -> (term * term) * T
+    Binding.T * term -> T -> (term * term) * T
   val revert_abbrev: string -> string -> T -> T
   val hide: bool -> string -> T -> T
   val empty: T
@@ -211,7 +211,7 @@
 fun err_dup_const c =
   error ("Duplicate declaration of constant " ^ quote c);
 
-fun extend_decls naming decl tab = NameSpace.table_declare naming decl tab
+fun extend_decls naming decl tab = NameSpace.bind naming decl tab
   handle Symtab.DUP c => err_dup_const c;
 
 
@@ -273,7 +273,7 @@
       |> cert_term;
     val normal_rhs = expand_term rhs;
     val T = Term.fastype_of rhs;
-    val lhs = Const (NameSpace.full_binding naming b, T);
+    val lhs = Const (NameSpace.full_name naming b, T);
 
     fun err msg = error (msg ^ " on rhs of abbreviation:\n" ^
       Pretty.string_of_term pp (Logic.mk_equals (lhs, rhs)));
--- a/src/Pure/context.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/context.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/context.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Generic theory contexts with unique identity, arbitrarily typed data,
@@ -392,14 +391,11 @@
 
 fun finish_thy thy = NAMED_CRITICAL "theory" (fn () =>
   let
-    val {name, version, intermediates} = history_of thy;
-    val rs = map ((fn TheoryRef r => r) o check_thy) intermediates;
-    val thy' as Theory ({self, id, ids, ...}, data', ancestry', _) = name_thy name thy;
-    val identity' = make_identity self id ids Inttab.empty;
+    val {name, ...} = history_of thy;
+    val Theory (identity', data', ancestry', _) = name_thy name thy;
     val history' = make_history name 0 [];
-    val thy'' = vitalize (Theory (identity', data', ancestry', history'));
-    val _ = List.app (fn r => r := thy'') rs;
-  in thy'' end);
+    val thy' = vitalize (Theory (identity', data', ancestry', history'));
+  in thy' end);
 
 
 (* theory data *)
--- a/src/Pure/facts.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/facts.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -31,9 +31,9 @@
   val props: T -> thm list
   val could_unify: T -> term -> thm list
   val merge: T * T -> T
-  val add_global: NameSpace.naming -> Name.binding * thm list -> T -> string * T
-  val add_local: bool -> NameSpace.naming -> Name.binding * thm list -> T -> string * T
-  val add_dynamic: NameSpace.naming -> Name.binding * (Context.generic -> thm list) -> T -> string * T
+  val add_global: NameSpace.naming -> Binding.T * thm list -> T -> string * T
+  val add_local: bool -> NameSpace.naming -> Binding.T * thm list -> T -> string * T
+  val add_dynamic: NameSpace.naming -> Binding.T * (Context.generic -> thm list) -> T -> string * T
   val del: string -> T -> T
   val hide: bool -> string -> T -> T
 end;
@@ -192,10 +192,10 @@
 
 fun add permissive do_props naming (b, ths) (Facts {facts, props}) =
   let
-    val (name, facts') = if Binding.is_nothing b then ("", facts)
+    val (name, facts') = if Binding.is_empty b then ("", facts)
       else let
         val (space, tab) = facts;
-        val (name, space') = NameSpace.declare_binding naming b space;
+        val (name, space') = NameSpace.declare naming b space;
         val entry = (name, (Static ths, serial ()));
         val tab' = (if permissive then Symtab.update else Symtab.update_new) entry tab
           handle Symtab.DUP dup => err_dup_fact dup;
@@ -213,7 +213,7 @@
 
 fun add_dynamic naming (b, f) (Facts {facts = (space, tab), props}) =
   let
-    val (name, space') = NameSpace.declare_binding naming b space;
+    val (name, space') = NameSpace.declare naming b space;
     val entry = (name, (Dynamic f, serial ()));
     val tab' = Symtab.update_new entry tab
       handle Symtab.DUP dup => err_dup_fact dup;
--- a/src/Pure/goal.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/goal.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -20,7 +20,7 @@
   val conclude: thm -> thm
   val finish: thm -> thm
   val norm_result: thm -> thm
-  val future_result: Proof.context -> (unit -> thm) -> term -> thm
+  val future_result: Proof.context -> thm future -> term -> thm
   val prove_internal: cterm list -> cterm -> (thm list -> tactic) -> thm
   val prove_multi: Proof.context -> string list -> term list -> term list ->
     ({prems: thm list, context: Proof.context} -> tactic) -> thm list
@@ -116,11 +116,11 @@
 
     val global_prop =
       Term.map_types Logic.varifyT (fold_rev Logic.all xs (Logic.list_implies (As, prop)));
-    val global_result =
-      Thm.generalize (map #1 tfrees, []) 0 o
-      Drule.forall_intr_list fixes o
-      Drule.implies_intr_list assms o
-      Thm.adjust_maxidx_thm ~1 o result;
+    val global_result = result |> Future.map
+      (Thm.adjust_maxidx_thm ~1 #>
+        Drule.implies_intr_list assms #>
+        Drule.forall_intr_list fixes #>
+        Thm.generalize (map #1 tfrees, []) 0);
     val local_result =
       Thm.future global_result (cert global_prop)
       |> Thm.instantiate (instT, [])
@@ -178,7 +178,7 @@
           end);
     val res =
       if immediate orelse #maxidx (Thm.rep_cterm stmt) >= 0 then result ()
-      else future_result ctxt' result (Thm.term_of stmt);
+      else future_result ctxt' (Future.fork_background result) (Thm.term_of stmt);
   in
     Conjunction.elim_balanced (length props) res
     |> map (Assumption.export false ctxt' ctxt)
--- a/src/Pure/more_thm.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/more_thm.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -281,7 +281,7 @@
   let
     val name' = if name = "" then "axiom_" ^ serial_string () else name;
     val thy' = thy |> Theory.add_axioms_i [(name', prop)];
-    val axm = unvarify (Thm.axiom thy' (Sign.full_name thy' name'));
+    val axm = unvarify (Thm.axiom thy' (Sign.full_bname thy' name'));
   in (axm, thy') end;
 
 fun add_def unchecked overloaded (name, prop) thy =
@@ -293,7 +293,7 @@
 
     val prop' = Term.map_types (Term.map_atyps (perhaps (AList.lookup (op =) strip_sorts))) prop;
     val thy' = Theory.add_defs_i unchecked overloaded [(name, prop')] thy;
-    val axm' = Thm.axiom thy' (Sign.full_name thy' name);
+    val axm' = Thm.axiom thy' (Sign.full_bname thy' name);
     val thm = unvarify (Thm.instantiate (recover_sorts, []) axm');
   in (thm, thy') end;
 
--- a/src/Pure/morphism.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/morphism.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/morphism.ML
-    ID:         $Id$
     Author:     Makarius
 
 Abstract morphisms on formal entities.
@@ -17,21 +16,21 @@
 signature MORPHISM =
 sig
   include BASIC_MORPHISM
-  val var: morphism -> Name.binding * mixfix -> Name.binding * mixfix
-  val name: morphism -> Name.binding -> Name.binding
+  val var: morphism -> Binding.T * mixfix -> Binding.T * mixfix
+  val binding: morphism -> Binding.T -> Binding.T
   val typ: morphism -> typ -> typ
   val term: morphism -> term -> term
   val fact: morphism -> thm list -> thm list
   val thm: morphism -> thm -> thm
   val cterm: morphism -> cterm -> cterm
   val morphism:
-   {name: Name.binding -> Name.binding,
-    var: Name.binding * mixfix -> Name.binding * mixfix,
+   {binding: Binding.T -> Binding.T,
+    var: Binding.T * mixfix -> Binding.T * mixfix,
     typ: typ -> typ,
     term: term -> term,
     fact: thm list -> thm list} -> morphism
-  val name_morphism: (Name.binding -> Name.binding) -> morphism
-  val var_morphism: (Name.binding * mixfix -> Name.binding * mixfix) -> morphism
+  val binding_morphism: (Binding.T -> Binding.T) -> morphism
+  val var_morphism: (Binding.T * mixfix -> Binding.T * mixfix) -> morphism
   val typ_morphism: (typ -> typ) -> morphism
   val term_morphism: (term -> term) -> morphism
   val fact_morphism: (thm list -> thm list) -> morphism
@@ -46,15 +45,15 @@
 struct
 
 datatype morphism = Morphism of
- {name: Name.binding -> Name.binding,
-  var: Name.binding * mixfix -> Name.binding * mixfix,
+ {binding: Binding.T -> Binding.T,
+  var: Binding.T * mixfix -> Binding.T * mixfix,
   typ: typ -> typ,
   term: term -> term,
   fact: thm list -> thm list};
 
 type declaration = morphism -> Context.generic -> Context.generic;
 
-fun name (Morphism {name, ...}) = name;
+fun binding (Morphism {binding, ...}) = binding;
 fun var (Morphism {var, ...}) = var;
 fun typ (Morphism {typ, ...}) = typ;
 fun term (Morphism {term, ...}) = term;
@@ -64,19 +63,19 @@
 
 val morphism = Morphism;
 
-fun name_morphism name = morphism {name = name, var = I, typ = I, term = I, fact = I};
-fun var_morphism var = morphism {name = I, var = var, typ = I, term = I, fact = I};
-fun typ_morphism typ = morphism {name = I, var = I, typ = typ, term = I, fact = I};
-fun term_morphism term = morphism {name = I, var = I, typ = I, term = term, fact = I};
-fun fact_morphism fact = morphism {name = I, var = I, typ = I, term = I, fact = fact};
-fun thm_morphism thm = morphism {name = I, var = I, typ = I, term = I, fact = map thm};
+fun binding_morphism binding = morphism {binding = binding, var = I, typ = I, term = I, fact = I};
+fun var_morphism var = morphism {binding = I, var = var, typ = I, term = I, fact = I};
+fun typ_morphism typ = morphism {binding = I, var = I, typ = typ, term = I, fact = I};
+fun term_morphism term = morphism {binding = I, var = I, typ = I, term = term, fact = I};
+fun fact_morphism fact = morphism {binding = I, var = I, typ = I, term = I, fact = fact};
+fun thm_morphism thm = morphism {binding = I, var = I, typ = I, term = I, fact = map thm};
 
-val identity = morphism {name = I, var = I, typ = I, term = I, fact = I};
+val identity = morphism {binding = I, var = I, typ = I, term = I, fact = I};
 
 fun compose
-    (Morphism {name = name1, var = var1, typ = typ1, term = term1, fact = fact1})
-    (Morphism {name = name2, var = var2, typ = typ2, term = term2, fact = fact2}) =
-  morphism {name = name1 o name2, var = var1 o var2,
+    (Morphism {binding = binding1, var = var1, typ = typ1, term = term1, fact = fact1})
+    (Morphism {binding = binding2, var = var2, typ = typ2, term = term2, fact = fact2}) =
+  morphism {binding = binding1 o binding2, var = var1 o var2,
     typ = typ1 o typ2, term = term1 o term2, fact = fact1 o fact2};
 
 fun phi1 $> phi2 = compose phi2 phi1;
--- a/src/Pure/name.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/name.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -1,8 +1,7 @@
 (*  Title:      Pure/name.ML
-    ID:         $Id$
     Author:     Makarius
 
-Names of basic logical entities (variables etc.).  Generic name bindings.
+Names of basic logical entities (variables etc.).
 *)
 
 signature NAME =
@@ -28,12 +27,6 @@
   val variants: string list -> context -> string list * context
   val variant_list: string list -> string list -> string list
   val variant: string list -> string -> string
-
-  include BINDING
-  type binding = Binding.T
-  val name_of: Binding.T -> string (*FIMXE legacy*)
-  val map_name: (string -> string) -> Binding.T -> Binding.T (*FIMXE legacy*)
-  val qualified: string -> Binding.T -> Binding.T (*FIMXE legacy*)
 end;
 
 structure Name: NAME =
@@ -145,16 +138,4 @@
 fun variant_list used names = #1 (make_context used |> variants names);
 fun variant used = singleton (variant_list used);
 
-
-(** generic name bindings **)
-
-open Binding;
-
-type binding = Binding.T;
-
-val prefix_of = fst o dest_binding;
-val name_of = snd o dest_binding;
-val map_name = map_binding o apsnd;
-val qualified = map_name o NameSpace.qualified;
-
 end;
--- a/src/Pure/proofterm.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/proofterm.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -22,10 +22,10 @@
    | PAxm of string * term * typ list option
    | Oracle of string * term * typ list option
    | Promise of serial * term * typ list
-   | PThm of serial * ((string * term * typ list option) * proof_body Lazy.T)
+   | PThm of serial * ((string * term * typ list option) * proof_body lazy)
   and proof_body = PBody of
     {oracles: (string * term) OrdList.T,
-     thms: (serial * (string * term * proof_body Lazy.T)) OrdList.T,
+     thms: (serial * (string * term * proof_body lazy)) OrdList.T,
      proof: proof}
 
   val %> : proof * term -> proof
@@ -36,10 +36,10 @@
   include BASIC_PROOFTERM
 
   type oracle = string * term
-  type pthm = serial * (string * term * proof_body Lazy.T)
-  val force_body: proof_body Lazy.T ->
+  type pthm = serial * (string * term * proof_body lazy)
+  val force_body: proof_body lazy ->
     {oracles: oracle OrdList.T, thms: pthm OrdList.T, proof: proof}
-  val force_proof: proof_body Lazy.T -> proof
+  val force_proof: proof_body lazy -> proof
   val proof_of: proof_body -> proof
   val fold_body_thms: (string * term * proof_body -> 'a -> 'a) -> proof_body list -> 'a -> 'a
   val fold_proof_atoms: bool -> (proof -> 'a -> 'a) -> proof list -> 'a -> 'a
@@ -111,7 +111,7 @@
   val promise_proof: theory -> serial -> term -> proof
   val fulfill_proof: theory -> (serial * proof) list -> proof_body -> proof_body
   val thm_proof: theory -> string -> term list -> term ->
-    (serial * proof) list Lazy.T -> proof_body -> pthm * proof
+    (serial * proof) list lazy -> proof_body -> pthm * proof
   val get_name: term list -> term -> proof -> string
 
   (** rewriting on proof terms **)
@@ -143,14 +143,14 @@
  | PAxm of string * term * typ list option
  | Oracle of string * term * typ list option
  | Promise of serial * term * typ list
- | PThm of serial * ((string * term * typ list option) * proof_body Lazy.T)
+ | PThm of serial * ((string * term * typ list option) * proof_body lazy)
 and proof_body = PBody of
   {oracles: (string * term) OrdList.T,
-   thms: (serial * (string * term * proof_body Lazy.T)) OrdList.T,
+   thms: (serial * (string * term * proof_body lazy)) OrdList.T,
    proof: proof};
 
 type oracle = string * term;
-type pthm = serial * (string * term * proof_body Lazy.T);
+type pthm = serial * (string * term * proof_body lazy);
 
 val force_body = Lazy.force #> (fn PBody args => args);
 val force_proof = #proof o force_body;
--- a/src/Pure/pure_thy.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/pure_thy.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -11,7 +11,7 @@
   val intern_fact: theory -> xstring -> string
   val defined_fact: theory -> string -> bool
   val hide_fact: bool -> string -> theory -> theory
-  val force_proofs: theory -> unit
+  val join_proofs: theory list -> unit Exn.result list
   val get_fact: Context.generic -> theory -> Facts.ref -> thm list
   val get_thms: theory -> xstring -> thm list
   val get_thm: theory -> xstring -> thm
@@ -31,9 +31,9 @@
   val add_thm: (bstring * thm) * attribute list -> theory -> thm * theory
   val add_thmss: ((bstring * thm list) * attribute list) list -> theory -> thm list list * theory
   val add_thms_dynamic: bstring * (Context.generic -> thm list) -> theory -> theory
-  val note_thmss: string -> ((Name.binding * attribute list) *
+  val note_thmss: string -> ((Binding.T * attribute list) *
     (thm list * attribute list) list) list -> theory -> (string * thm list) list * theory
-  val note_thmss_grouped: string -> string -> ((Name.binding * attribute list) *
+  val note_thmss_grouped: string -> string -> ((Binding.T * attribute list) *
     (thm list * attribute list) list) list -> theory -> (string * thm list) list * theory
   val add_axioms: ((bstring * term) * attribute list) list -> theory -> thm list * theory
   val add_axioms_cmd: ((bstring * string) * attribute list) list -> theory -> thm list * theory
@@ -59,7 +59,7 @@
 
 structure FactsData = TheoryDataFun
 (
-  type T = Facts.T * unit Lazy.T list;
+  type T = Facts.T * unit lazy list;
   val empty = (Facts.empty, []);
   val copy = I;
   fun extend (facts, _) = (facts, []);
@@ -79,8 +79,9 @@
 
 (* proofs *)
 
-fun lazy_proof thm = Lazy.lazy (fn () => Thm.force_proof thm);
-val force_proofs = List.app Lazy.force o #2 o FactsData.get;
+fun lazy_proof thm = Lazy.lazy (fn () => Thm.join_proof thm);
+
+val join_proofs = maps (fn thy => map (Exn.capture Lazy.force) (rev (#2 (FactsData.get thy))));
 
 
 
@@ -144,11 +145,11 @@
   (FactsData.map (apsnd (fold (cons o lazy_proof) thms)) thy, thms);
 
 fun enter_thms pre_name post_name app_att (b, thms) thy =
-  if Binding.is_nothing b
+  if Binding.is_empty b
   then swap (enter_proofs (app_att (thy, thms)))
   else let
     val naming = Sign.naming_of thy;
-    val name = NameSpace.full_binding naming b;
+    val name = NameSpace.full_name naming b;
     val (thy', thms') =
       enter_proofs (apsnd (post_name name) (app_att (thy, pre_name name thms)));
     val thms'' = map (Thm.transfer thy') thms';
@@ -159,20 +160,20 @@
 (* store_thm(s) *)
 
 fun store_thms (bname, thms) = enter_thms (name_thms true true Position.none)
-  (name_thms false true Position.none) I (Name.binding bname, thms);
+  (name_thms false true Position.none) I (Binding.name bname, thms);
 
 fun store_thm (bname, th) = store_thms (bname, [th]) #>> the_single;
 
 fun store_thm_open (bname, th) =
   enter_thms (name_thms true false Position.none) (name_thms false false Position.none) I
-    (Name.binding bname, [th]) #>> the_single;
+    (Binding.name bname, [th]) #>> the_single;
 
 
 (* add_thms(s) *)
 
 fun add_thms_atts pre_name ((bname, thms), atts) =
   enter_thms pre_name (name_thms false true Position.none)
-    (foldl_map (Thm.theory_attributes atts)) (Name.binding bname, thms);
+    (foldl_map (Thm.theory_attributes atts)) (Binding.name bname, thms);
 
 fun gen_add_thmss pre_name =
   fold_map (add_thms_atts pre_name);
@@ -189,7 +190,7 @@
 
 fun add_thms_dynamic (bname, f) thy = thy
   |> (FactsData.map o apfst)
-      (Facts.add_dynamic (Sign.naming_of thy) (Name.binding bname, f) #> snd);
+      (Facts.add_dynamic (Sign.naming_of thy) (Binding.name bname, f) #> snd);
 
 
 (* note_thmss *)
@@ -199,7 +200,7 @@
 fun gen_note_thmss tag = fold_map (fn ((b, more_atts), ths_atts) => fn thy =>
   let
     val pos = Binding.pos_of b;
-    val name = Sign.full_binding thy b;
+    val name = Sign.full_name thy b;
     val _ = Position.report (Markup.fact_decl name) pos;
 
     fun app (x, (ths, atts)) = foldl_map (Thm.theory_attributes atts) (x, ths);
@@ -219,7 +220,7 @@
 (* store axioms as theorems *)
 
 local
-  fun get_ax thy (name, _) = Thm.axiom thy (Sign.full_name thy name);
+  fun get_ax thy (name, _) = Thm.axiom thy (Sign.full_bname thy name);
   fun get_axs thy named_axs = map (Thm.forall_elim_vars 0 o get_ax thy) named_axs;
   fun add_axm add = fold_map (fn ((name, ax), atts) => fn thy =>
     let
--- a/src/Pure/sign.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/sign.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -14,11 +14,10 @@
     tsig: Type.tsig,
     consts: Consts.T}
   val naming_of: theory -> NameSpace.naming
+  val full_name: theory -> Binding.T -> string
   val base_name: string -> bstring
-  val full_name: theory -> bstring -> string
-  val full_binding: theory -> Name.binding -> string
-  val full_name_path: theory -> string -> bstring -> string
-  val declare_name: theory -> string -> NameSpace.T -> NameSpace.T
+  val full_bname: theory -> bstring -> string
+  val full_bname_path: theory -> string -> bstring -> string
   val syn_of: theory -> Syntax.syntax
   val tsig_of: theory -> Type.tsig
   val classes_of: theory -> Sorts.algebra
@@ -92,10 +91,10 @@
   val del_modesyntax: Syntax.mode -> (bstring * string * mixfix) list -> theory -> theory
   val del_modesyntax_i: Syntax.mode -> (bstring * typ * mixfix) list -> theory -> theory
   val notation: bool -> Syntax.mode -> (term * mixfix) list -> theory -> theory
-  val declare_const: Properties.T -> (Name.binding * typ) * mixfix -> theory -> term * theory
+  val declare_const: Properties.T -> (Binding.T * typ) * mixfix -> theory -> term * theory
   val add_consts: (bstring * string * mixfix) list -> theory -> theory
   val add_consts_i: (bstring * typ * mixfix) list -> theory -> theory
-  val add_abbrev: string -> Properties.T -> Name.binding * term -> theory -> (term * term) * theory
+  val add_abbrev: string -> Properties.T -> Binding.T * term -> theory -> (term * term) * theory
   val revert_abbrev: string -> string -> theory -> theory
   val add_const_constraint: string * typ option -> theory -> theory
   val primitive_class: string * class list -> theory -> theory
@@ -189,10 +188,10 @@
 val naming_of = #naming o rep_sg;
 val base_name = NameSpace.base;
 
-val full_name = NameSpace.full o naming_of;
-val full_binding = NameSpace.full_binding o naming_of;
-fun full_name_path thy elems = NameSpace.full (NameSpace.add_path elems (naming_of thy));
-val declare_name = NameSpace.declare o naming_of;
+val full_name = NameSpace.full_name o naming_of;
+fun full_bname thy = NameSpace.full_name (naming_of thy) o Binding.name;
+fun full_bname_path thy elems =
+  NameSpace.full_name (NameSpace.add_path elems (naming_of thy)) o Binding.name;
 
 
 (* syntax *)
@@ -509,12 +508,12 @@
     val prepT = Type.no_tvars o Term.no_dummyT o certify_typ thy o parse_typ ctxt;
     fun prep (raw_b, raw_T, raw_mx) =
       let
-        val (mx_name, mx) = Syntax.const_mixfix (Name.name_of raw_b) raw_mx;
-        val b = Name.map_name (K mx_name) raw_b;
-        val c = full_binding thy b;
-        val c_syn = if authentic then Syntax.constN ^ c else Name.name_of b;
+        val (mx_name, mx) = Syntax.const_mixfix (Binding.base_name raw_b) raw_mx;
+        val b = Binding.map_base (K mx_name) raw_b;
+        val c = full_name thy b;
+        val c_syn = if authentic then Syntax.constN ^ c else Binding.base_name b;
         val T = (prepT raw_T handle TYPE (msg, _, _) => error msg) handle ERROR msg =>
-          cat_error msg ("in declaration of constant " ^ quote (Name.display b));
+          cat_error msg ("in declaration of constant " ^ quote (Binding.display b));
         val T' = Logic.varifyT T;
       in ((b, T'), (c_syn, T', mx), Const (c, T)) end;
     val args = map prep raw_args;
@@ -526,7 +525,7 @@
     |> pair (map #3 args)
   end;
 
-fun bindify (name, T, mx) = (Name.binding name, T, mx);
+fun bindify (name, T, mx) = (Binding.name name, T, mx);
 
 in
 
@@ -551,7 +550,7 @@
     val pp = Syntax.pp_global thy;
     val prep_tm = no_frees pp o Term.no_dummy_patterns o cert_term_abbrev thy;
     val t = (prep_tm raw_t handle TYPE (msg, _, _) => error msg | TERM (msg, _) => error msg)
-      handle ERROR msg => cat_error msg ("in constant abbreviation " ^ quote (Name.display b));
+      handle ERROR msg => cat_error msg ("in constant abbreviation " ^ quote (Binding.display b));
     val (res, consts') = consts_of thy
       |> Consts.abbreviate pp (tsig_of thy) (naming_of thy) mode tags (b, t);
   in (res, thy |> map_consts (K consts')) end;
--- a/src/Pure/simplifier.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/simplifier.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -177,9 +177,10 @@
 
 fun gen_simproc prep {name, lhss, proc, identifier} lthy =
   let
+    val b = Binding.name name;
     val naming = LocalTheory.full_naming lthy;
     val simproc = make_simproc
-      {name = LocalTheory.full_name lthy name,
+      {name = LocalTheory.full_name lthy b,
        lhss =
         let
           val lhss' = prep lthy lhss;
@@ -194,11 +195,11 @@
   in
     lthy |> LocalTheory.declaration (fn phi =>
       let
-        val b' = Morphism.name phi (Name.binding name);
+        val b' = Morphism.binding phi b;
         val simproc' = morph_simproc phi simproc;
       in
         Simprocs.map (fn simprocs =>
-          NameSpace.table_declare naming (b', simproc') simprocs |> snd
+          NameSpace.bind naming (b', simproc') simprocs |> snd
             handle Symtab.DUP dup => err_dup_simproc dup)
         #> map_ss (fn ss => ss addsimprocs [simproc'])
       end)
--- a/src/Pure/theory.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/theory.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -36,7 +36,7 @@
   val add_defs_i: bool -> bool -> (bstring * term) list -> theory -> theory
   val add_finals: bool -> string list -> theory -> theory
   val add_finals_i: bool -> term list -> theory -> theory
-  val specify_const: Properties.T -> (Name.binding * typ) * mixfix -> theory -> term * theory
+  val specify_const: Properties.T -> (Binding.T * typ) * mixfix -> theory -> term * theory
 end
 
 structure Theory: THEORY =
@@ -178,8 +178,8 @@
 
 fun gen_add_axioms prep_axm raw_axms thy = thy |> map_axioms (fn axioms =>
   let
-    val axms = map (apsnd Logic.varify o prep_axm thy) raw_axms;
-    val axioms' = NameSpace.extend_table (Sign.naming_of thy) axms axioms
+    val axms = map (apfst Binding.name o apsnd Logic.varify o prep_axm thy) raw_axms;
+    val axioms' = fold (snd oo NameSpace.bind (Sign.naming_of thy)) axms axioms
       handle Symtab.DUP dup => err_dup_axm dup;
   in axioms' end);
 
@@ -253,7 +253,7 @@
 fun check_def thy unchecked overloaded (bname, tm) defs =
   let
     val ctxt = ProofContext.init thy;
-    val name = Sign.full_name thy bname;
+    val name = Sign.full_bname thy bname;
     val (lhs_const, rhs) = Sign.cert_def ctxt tm;
     val rhs_consts = fold_aterms (fn Const const => insert (op =) const | _ => I) rhs [];
     val _ = check_overloading thy overloaded lhs_const;
--- a/src/Pure/thm.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/thm.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -146,11 +146,10 @@
   val varifyT: thm -> thm
   val varifyT': (string * sort) list -> thm -> ((string * sort) * indexname) list * thm
   val freezeT: thm -> thm
-  val join_futures: theory -> unit
-  val future: (unit -> thm) -> cterm -> thm
+  val future: thm future -> cterm -> thm
   val proof_body_of: thm -> proof_body
   val proof_of: thm -> proof
-  val force_proof: thm -> unit
+  val join_proof: thm -> unit
   val extern_oracles: theory -> xstring list
   val add_oracle: bstring * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory
 end;
@@ -347,8 +346,9 @@
   tpairs: (term * term) list,                   (*flex-flex pairs*)
   prop: term}                                   (*conclusion*)
 and deriv = Deriv of
- {all_promises: (serial * thm Future.T) OrdList.T,
-  promises: (serial * thm Future.T) OrdList.T,
+ {max_promise: serial,
+  open_promises: (serial * thm future) OrdList.T,
+  promises: (serial * thm future) OrdList.T,
   body: Pt.proof_body};
 
 type conv = cterm -> thm;
@@ -502,12 +502,11 @@
 
 (** derivations **)
 
-fun make_deriv all_promises promises oracles thms proof =
-  Deriv {all_promises = all_promises, promises = promises,
+fun make_deriv max_promise open_promises promises oracles thms proof =
+  Deriv {max_promise = max_promise, open_promises = open_promises, promises = promises,
     body = PBody {oracles = oracles, thms = thms, proof = proof}};
 
-val closed_deriv = make_deriv [] [] [] [];
-val empty_deriv = closed_deriv Pt.MinProof;
+val empty_deriv = make_deriv ~1 [] [] [] [] Pt.MinProof;
 
 
 (* inference rules *)
@@ -515,12 +514,13 @@
 fun promise_ord ((i, _), (j, _)) = int_ord (j, i);
 
 fun deriv_rule2 f
-    (Deriv {all_promises = all_ps1, promises = ps1,
+    (Deriv {max_promise = max1, open_promises = open_ps1, promises = ps1,
       body = PBody {oracles = oras1, thms = thms1, proof = prf1}})
-    (Deriv {all_promises = all_ps2, promises = ps2,
+    (Deriv {max_promise = max2, open_promises = open_ps2, promises = ps2,
       body = PBody {oracles = oras2, thms = thms2, proof = prf2}}) =
   let
-    val all_ps = OrdList.union promise_ord all_ps1 all_ps2;
+    val max = Int.max (max1, max2);
+    val open_ps = OrdList.union promise_ord open_ps1 open_ps2;
     val ps = OrdList.union promise_ord ps1 ps2;
     val oras = Pt.merge_oracles oras1 oras2;
     val thms = Pt.merge_thms thms1 thms2;
@@ -530,10 +530,10 @@
       | 1 => MinProof
       | 0 => MinProof
       | i => error ("Illegal level of detail for proof objects: " ^ string_of_int i));
-  in make_deriv all_ps ps oras thms prf end;
+  in make_deriv max open_ps ps oras thms prf end;
 
 fun deriv_rule1 f = deriv_rule2 (K f) empty_deriv;
-fun deriv_rule0 prf = deriv_rule1 I (closed_deriv prf);
+fun deriv_rule0 prf = deriv_rule1 I (make_deriv ~1 [] [] [] [] prf);
 
 
 
@@ -1587,36 +1587,7 @@
 
 
 
-(*** Promises ***)
-
-(* pending future derivations *)
-
-structure Futures = TheoryDataFun
-(
-  type T = thm Future.T list ref;
-  val empty : T = ref [];
-  val copy = I;  (*shared ref within all versions of whole theory body*)
-  fun extend _ : T = ref [];
-  fun merge _ _ : T = ref [];
-);
-
-val _ = Context.>> (Context.map_theory Futures.init);
-
-fun add_future thy future = CRITICAL (fn () => change (Futures.get thy) (cons future));
-
-fun join_futures thy =
-  let
-    val futures = Futures.get thy;
-    fun joined () =
-     (List.app (ignore o Future.join_result) (rev (! futures));
-      CRITICAL (fn () =>
-        let
-          val (finished, unfinished) = List.partition Future.is_finished (! futures);
-          val _ = futures := unfinished;
-        in finished end)
-      |> Future.join_results |> Exn.release_all |> null);
-  in while not (joined ()) do () end;
-
+(*** Future theorems -- proofs with promises ***)
 
 (* future rule *)
 
@@ -1627,26 +1598,25 @@
     val _ = Theory.check_thy orig_thy;
     fun err msg = raise THM ("future_result: " ^ msg, 0, [thm]);
 
-    val Thm (Deriv {all_promises, ...}, {shyps, hyps, tpairs, prop, ...}) = thm;
+    val Thm (Deriv {max_promise, ...}, {shyps, hyps, tpairs, prop, ...}) = thm;
     val _ = prop aconv orig_prop orelse err "bad prop";
     val _ = null tpairs orelse err "bad tpairs";
     val _ = null hyps orelse err "bad hyps";
     val _ = Sorts.subset (shyps, orig_shyps) orelse err "bad shyps";
-    val _ = forall (fn (j, _) => j < i) all_promises orelse err "bad dependencies";
+    val _ = max_promise < i orelse err "bad dependencies";
   in thm end;
 
-fun future make_result ct =
+fun future future_thm ct =
   let
     val Cterm {thy_ref = thy_ref, t = prop, T, maxidx, sorts} = ct;
     val thy = Context.reject_draft (Theory.deref thy_ref);
     val _ = T <> propT andalso raise CTERM ("future: prop expected", [ct]);
 
     val i = serial ();
-    val future = Future.fork_background (future_result i thy sorts prop o norm_proof o make_result);
-    val _ = add_future thy future;
+    val future = future_thm |> Future.map (future_result i thy sorts prop o norm_proof);
     val promise = (i, future);
   in
-    Thm (make_deriv [promise] [promise] [] [] (Pt.promise_proof thy i prop),
+    Thm (make_deriv i [promise] [promise] [] [] (Pt.promise_proof thy i prop),
      {thy_ref = thy_ref,
       tags = [],
       maxidx = maxidx,
@@ -1661,14 +1631,14 @@
 
 fun raw_proof_of (Thm (Deriv {body, ...}, _)) = Proofterm.proof_of body;
 
-fun proof_body_of (Thm (Deriv {all_promises, promises, body}, {thy_ref, ...})) =
+fun proof_body_of (Thm (Deriv {open_promises, promises, body, ...}, {thy_ref, ...})) =
   let
-    val _ = Exn.release_all (map (Future.join_result o #2) (rev all_promises));
+    val _ = Exn.release_all (map (Future.join_result o #2) (rev open_promises));
     val ps = map (apsnd (raw_proof_of o Future.join)) promises;
   in Pt.fulfill_proof (Theory.deref thy_ref) ps body end;
 
 val proof_of = Proofterm.proof_of o proof_body_of;
-val force_proof = ignore o proof_body_of;
+val join_proof = ignore o proof_body_of;
 
 
 (* closed derivations with official name *)
@@ -1678,14 +1648,17 @@
 
 fun put_name name (thm as Thm (der, args)) =
   let
-    val Deriv {promises, body, ...} = der;
+    val Deriv {max_promise, open_promises, promises, body, ...} = der;
     val {thy_ref, hyps, prop, tpairs, ...} = args;
-    val _ = null tpairs orelse raise THM ("name_thm: unsolved flex-flex constraints", 0, [thm]);
+    val _ = null tpairs orelse raise THM ("put_name: unsolved flex-flex constraints", 0, [thm]);
 
     val ps = Lazy.lazy (fn () => map (apsnd (proof_of o Future.join)) promises);
     val thy = Theory.deref thy_ref;
     val (pthm, proof) = Pt.thm_proof thy name hyps prop ps body;
-    val der' = make_deriv [] [] [] [pthm] proof;
+
+    val open_promises' = open_promises |> filter (fn (_, p) =>
+      (case Future.peek p of SOME (Exn.Result _) => false | _ => true));
+    val der' = make_deriv max_promise open_promises' [] [] [pthm] proof;
     val _ = Theory.check_thy thy;
   in Thm (der', args) end;
 
@@ -1701,7 +1674,7 @@
       raise THM ("Oracle's result must have type prop: " ^ name, 0, [])
     else
       let val prf = Pt.oracle_proof name prop in
-        Thm (make_deriv [] [] (Pt.make_oracles prf) [] prf,
+        Thm (make_deriv ~1 [] [] (Pt.make_oracles prf) [] prf,
          {thy_ref = Theory.merge_refs (thy_ref1, thy_ref2),
           tags = [],
           maxidx = maxidx,
@@ -1732,9 +1705,9 @@
 fun add_oracle (bname, oracle) thy =
   let
     val naming = Sign.naming_of thy;
-    val name = NameSpace.full naming bname;
+    val name = NameSpace.full_name naming (Binding.name bname);
     val thy' = thy |> Oracles.map (fn (space, tab) =>
-      (NameSpace.declare naming name space,
+      (NameSpace.declare naming (Binding.name bname) space |> snd,
         Symtab.update_new (name, stamp ()) tab handle Symtab.DUP dup => err_dup_ora dup));
   in ((name, invoke_oracle (Theory.check_thy thy') name oracle), thy') end;
 
--- a/src/Pure/type.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/type.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -509,10 +509,9 @@
 fun add_class pp naming (c, cs) tsig =
   tsig |> map_tsig (fn ((space, classes), default, types) =>
     let
-      val c' = NameSpace.full naming c;
       val cs' = map (cert_class tsig) cs
         handle TYPE (msg, _, _) => error msg;
-      val space' = space |> NameSpace.declare naming c';
+      val (c', space') = space |> NameSpace.declare naming (Binding.name c);
       val classes' = classes |> Sorts.add_class pp (c', cs');
     in ((space', classes'), default, types) end);
 
@@ -568,8 +567,7 @@
 fun new_decl naming tags (c, decl) (space, types) =
   let
     val tags' = tags |> Position.default_properties (Position.thread_data ());
-    val c' = NameSpace.full naming c;
-    val space' = NameSpace.declare naming c' space;
+    val (c', space') = NameSpace.declare naming (Binding.name c) space;
     val types' =
       (case Symtab.lookup types c' of
         SOME ((decl', _), _) => err_in_decls c' decl decl'
--- a/src/Pure/variable.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Pure/variable.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -398,7 +398,7 @@
     val fact = export inner outer;
     val term = singleton (export_terms inner outer);
     val typ = Logic.type_map term;
-  in Morphism.morphism {name = I, var = I, typ = typ, term = term, fact = fact} end;
+  in Morphism.morphism {binding = I, var = I, typ = typ, term = term, fact = fact} end;
 
 
 
--- a/src/Tools/code/code_ml.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Tools/code/code_ml.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -912,7 +912,7 @@
 
 structure CodeAntiqData = ProofDataFun
 (
-  type T = string list * (bool * (string * (string * (string * string) list) Lazy.T));
+  type T = string list * (bool * (string * (string * (string * string) list) lazy));
   fun init _ = ([], (true, ("", Lazy.value ("", []))));
 );
 
--- a/src/Tools/induct.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/Tools/induct.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Tools/induct.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Proof by cases, induction, and coinduction.
@@ -51,7 +50,7 @@
   val setN: string
   (*proof methods*)
   val fix_tac: Proof.context -> int -> (string * typ) list -> int -> tactic
-  val add_defs: (Name.binding option * term) option list -> Proof.context ->
+  val add_defs: (Binding.T option * term) option list -> Proof.context ->
     (term option list * thm list) * Proof.context
   val atomize_term: theory -> term -> term
   val atomize_tac: int -> tactic
@@ -63,7 +62,7 @@
   val cases_tac: Proof.context -> term option list list -> thm option ->
     thm list -> int -> cases_tactic
   val get_inductT: Proof.context -> term option list list -> thm list list
-  val induct_tac: Proof.context -> (Name.binding option * term) option list list ->
+  val induct_tac: Proof.context -> (Binding.T option * term) option list list ->
     (string * typ) list list -> term option list -> thm list option ->
     thm list -> int -> cases_tactic
   val coinduct_tac: Proof.context -> term option list -> term option list -> thm option ->
@@ -553,7 +552,7 @@
   let
     fun add (SOME (SOME x, t)) ctxt =
           let val ([(lhs, (_, th))], ctxt') =
-            LocalDefs.add_defs [((x, NoSyn), ((Name.no_binding, []), t))] ctxt
+            LocalDefs.add_defs [((x, NoSyn), ((Binding.empty, []), t))] ctxt
           in ((SOME lhs, [th]), ctxt') end
       | add (SOME (NONE, t)) ctxt = ((SOME t, []), ctxt)
       | add NONE ctxt = ((NONE, []), ctxt);
--- a/src/ZF/Tools/datatype_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/ZF/Tools/datatype_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -101,7 +101,7 @@
   val npart = length rec_names;  (*number of mutually recursive parts*)
 
 
-  val full_name = Sign.full_name thy_path;
+  val full_name = Sign.full_bname thy_path;
 
   (*Make constructor definition;
     kpart is the number of this mutually recursive part*)
@@ -262,7 +262,7 @@
              ||> add_recursor
              ||> Sign.parent_path
 
-  val intr_names = map (Name.binding o #2) (List.concat con_ty_lists);
+  val intr_names = map (Binding.name o #2) (List.concat con_ty_lists);
   val (thy1, ind_result) =
     thy0 |> Ind_Package.add_inductive_i
       false (rec_tms, dom_sum) (map Thm.no_attributes (intr_names ~~ intr_tms))
--- a/src/ZF/Tools/inductive_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/ZF/Tools/inductive_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -29,10 +29,10 @@
   (*Insert definitions for the recursive sets, which
      must *already* be declared as constants in parent theory!*)
   val add_inductive_i: bool -> term list * term ->
-    ((Name.binding * term) * attribute list) list ->
+    ((Binding.T * term) * attribute list) list ->
     thm list * thm list * thm list * thm list -> theory -> theory * inductive_result
   val add_inductive: string list * string ->
-    ((Name.binding * string) * Attrib.src list) list ->
+    ((Binding.T * string) * Attrib.src list) list ->
     (Facts.ref * Attrib.src list) list * (Facts.ref * Attrib.src list) list *
     (Facts.ref * Attrib.src list) list * (Facts.ref * Attrib.src list) list ->
     theory -> theory * inductive_result
@@ -67,7 +67,7 @@
   val _ = Theory.requires thy "Inductive_ZF" "(co)inductive definitions";
   val ctxt = ProofContext.init thy;
 
-  val intr_specs = map (apfst (apfst Name.name_of)) raw_intr_specs;
+  val intr_specs = map (apfst (apfst Binding.base_name)) raw_intr_specs;
   val (intr_names, intr_tms) = split_list (map fst intr_specs);
   val case_names = RuleCases.case_names intr_names;
 
--- a/src/ZF/Tools/primrec_package.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/ZF/Tools/primrec_package.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -9,8 +9,8 @@
 
 signature PRIMREC_PACKAGE =
 sig
-  val add_primrec: ((Name.binding * string) * Attrib.src list) list -> theory -> theory * thm list
-  val add_primrec_i: ((Name.binding * term) * attribute list) list -> theory -> theory * thm list
+  val add_primrec: ((Binding.T * string) * Attrib.src list) list -> theory -> theory * thm list
+  val add_primrec_i: ((Binding.T * term) * attribute list) list -> theory -> theory * thm list
 end;
 
 structure PrimrecPackage : PRIMREC_PACKAGE =
@@ -180,7 +180,7 @@
 
     val (eqn_thms', thy2) =
       thy1
-      |> PureThy.add_thms ((map Name.name_of eqn_names ~~ eqn_thms) ~~ eqn_atts);
+      |> PureThy.add_thms ((map Binding.base_name eqn_names ~~ eqn_thms) ~~ eqn_atts);
     val (_, thy3) =
       thy2
       |> PureThy.add_thmss [(("simps", eqn_thms'), [Simplifier.simp_add])]
--- a/src/ZF/ind_syntax.ML	Wed Dec 10 10:23:47 2008 +0100
+++ b/src/ZF/ind_syntax.ML	Wed Dec 10 10:28:16 2008 +0100
@@ -85,7 +85,7 @@
       Logic.list_implies
         (map FOLogic.mk_Trueprop prems,
 	 FOLogic.mk_Trueprop
-	    (@{const mem} $ list_comb (Const (Sign.full_name sg name, T), args)
+	    (@{const mem} $ list_comb (Const (Sign.full_bname sg name, T), args)
 	               $ rec_tm))
   in  map mk_intr constructs  end;