Merged.
authorballarin
Sat, 28 Mar 2009 00:13:01 +0100
changeset 30752 5272864d6892
parent 30751 36a255c2e428 (current diff)
parent 30748 fe67d729a61c (diff)
child 30753 78d12065c638
child 30779 492ca5d4f235
Merged.
doc-src/Codegen/codegen_process.pdf
doc-src/Codegen/codegen_process.ps
src/FOL/ex/LocaleTest.thy
src/HOL/Arith_Tools.thy
src/HOL/ex/Code_Antiq.thy
src/HOL/ex/ImperativeQuicksort.thy
src/HOL/ex/Subarray.thy
src/HOL/ex/Sublist.thy
src/Pure/Isar/antiquote.ML
--- a/Admin/Mercurial/cvsids	Sat Mar 28 00:11:02 2009 +0100
+++ b/Admin/Mercurial/cvsids	Sat Mar 28 00:13:01 2009 +0100
@@ -1,6 +1,6 @@
 Identifiers of some old CVS file versions
 =========================================
 
+src/Pure/General/file.ML    1.18    6cdd6a8da9b9
+src/Pure/thm.ML             1.189   4b339d3907a0  (referenced in 25f28f9c28a3 as "2005-01-24 (revision 1.44)")
 src/Pure/type.ML            1.65    0d984ee030a1
-src/Pure/General/file.ML    1.18    6cdd6a8da9b9
-
--- a/Admin/isatest/isatest-stats	Sat Mar 28 00:11:02 2009 +0100
+++ b/Admin/isatest/isatest-stats	Sat Mar 28 00:13:01 2009 +0100
@@ -1,13 +1,12 @@
 #!/usr/bin/env bash
 #
-# $Id$
 # Author: Makarius
 #
 # DESCRIPTION: Standard statistics.
 
 THIS=$(cd "$(dirname "$0")"; pwd -P)
 
-PLATFORMS="at-poly at-sml-dev at64-poly at-poly-5.1-para-e at64-poly-5.1-para at-mac-poly-5.1-para afp"
+PLATFORMS="at-poly at64-poly at-poly-5.1-para-e at64-poly-5.1-para at-mac-poly-5.1-para afp at-sml-dev"
 
 ISABELLE_SESSIONS="\
   HOL-Plain \
@@ -31,6 +30,8 @@
   HOL-UNITY \
   HOL-Word \
   HOL-ex \
+  HOLCF \
+  IOA \
   ZF \
   ZF-Constructible \
   ZF-UNITY"
--- a/Admin/isatest/settings/at-mac-poly-5.1-para	Sat Mar 28 00:11:02 2009 +0100
+++ b/Admin/isatest/settings/at-mac-poly-5.1-para	Sat Mar 28 00:13:01 2009 +0100
@@ -4,7 +4,7 @@
   ML_SYSTEM="polyml-5.2.1"
   ML_PLATFORM="x86-darwin"
   ML_HOME="$POLYML_HOME/$ML_PLATFORM"
-  ML_OPTIONS="--mutable 200 --immutable 800"
+  ML_OPTIONS="--mutable 800 --immutable 2000"
 
 
 ISABELLE_HOME_USER=~/isabelle-at-mac-poly-e
--- a/NEWS	Sat Mar 28 00:11:02 2009 +0100
+++ b/NEWS	Sat Mar 28 00:13:01 2009 +0100
@@ -139,8 +139,8 @@
 INCOMPATBILITY.
 
 * Complete re-implementation of locales.  INCOMPATIBILITY.  The most
-important changes are listed below.  See documentation (forthcoming)
-and tutorial (also forthcoming) for details.
+important changes are listed below.  See the Tutorial on Locales for
+details.
 
 - In locale expressions, instantiation replaces renaming.  Parameters
 must be declared in a for clause.  To aid compatibility with previous
@@ -154,19 +154,23 @@
 
 - More flexible mechanisms to qualify names generated by locale
 expressions.  Qualifiers (prefixes) may be specified in locale
-expressions.  Available are normal qualifiers (syntax "name:") and
-strict qualifiers (syntax "name!:").  The latter must occur in name
-references and are useful to avoid accidental hiding of names, the
-former are optional.  Qualifiers derived from the parameter names of a
-locale are no longer generated.
+expressions, and can be marked as mandatory (syntax: "name!:") or
+optional (syntax "name?:").  The default depends for plain "name:"
+depends on the situation where a locale expression is used: in
+commands 'locale' and 'sublocale' prefixes are optional, in
+'interpretation' and 'interpret' prefixes are mandatory.  Old-style
+implicit qualifiers derived from the parameter names of a locale are
+no longer generated.
 
 - "sublocale l < e" replaces "interpretation l < e".  The
 instantiation clause in "interpretation" and "interpret" (square
 brackets) is no longer available.  Use locale expressions.
 
-- When converting proof scripts, be sure to replace qualifiers in
-"interpretation" and "interpret" by strict qualifiers.  Qualifiers in
-locale expressions range over a single locale instance only.
+- When converting proof scripts, be sure to mandatory qualifiers in
+'interpretation' and 'interpret' should be retained by default, even
+if this is an INCOMPATIBILITY compared to former behaviour.
+Qualifiers in locale expressions range over a single locale instance
+only.
 
 * Command 'instance': attached definitions no longer accepted.
 INCOMPATIBILITY, use proper 'instantiation' target.
@@ -176,30 +180,28 @@
 * The 'axiomatization' command now only works within a global theory
 context.  INCOMPATIBILITY.
 
-* New find_theorems criterion "solves" matching theorems that directly
-solve the current goal. Try "find_theorems solves".
-
-* Added an auto solve option, which can be enabled through the
-ProofGeneral Isabelle settings menu (disabled by default).
- 
-When enabled, find_theorems solves is called whenever a new lemma is
-stated. Any theorems that could solve the lemma directly are listed
-underneath the goal.
-
-* New command 'find_consts' searches for constants based on type and
-name patterns, e.g.
+* New 'find_theorems' criterion "solves" matching theorems that
+directly solve the current goal.
+
+* Auto solve feature for main theorem statements (cf. option in Proof
+General Isabelle settings menu, disabled by default).  Whenever a new
+goal is stated, "find_theorems solves" is called; any theorems that
+could solve the lemma directly are listed as part of the goal state.
+
+* Command 'find_consts' searches for constants based on type and name
+patterns, e.g.
 
     find_consts "_ => bool"
 
 By default, matching is against subtypes, but it may be restricted to
-the whole type. Searching by name is possible. Multiple queries are
+the whole type.  Searching by name is possible.  Multiple queries are
 conjunctive and queries may be negated by prefixing them with a
 hyphen:
 
     find_consts strict: "_ => bool" name: "Int" -"int => int"
 
-* New command 'local_setup' is similar to 'setup', but operates on a
-local theory context.
+* Command 'local_setup' is similar to 'setup', but operates on a local
+theory context.
 
 
 *** Document preparation ***
@@ -332,6 +334,11 @@
 * Simplifier: simproc for let expressions now unfolds if bound variable
 occurs at most once in let expression body.  INCOMPATIBILITY.
 
+* New attribute "arith" for facts that should always be used automaticaly
+by arithmetic. It is intended to be used locally in proofs, eg
+assumes [arith]: "x > 0"
+Global usage is discouraged because of possible performance impact.
+
 * New classes "top" and "bot" with corresponding operations "top" and "bot"
 in theory Orderings;  instantiation of class "complete_lattice" requires
 instantiation of classes "top" and "bot".  INCOMPATIBILITY.
@@ -495,10 +502,9 @@
 resulting ML value/function/datatype constructor binding in place.
 All occurrences of @{code} with a single ML block are generated
 simultaneously.  Provides a generic and safe interface for
-instrumentalizing code generation.  See HOL/ex/Code_Antiq for a toy
-example, or HOL/Complex/ex/ReflectedFerrack for a more ambitious
-application.  In future you ought refrain from ad-hoc compiling
-generated SML code on the ML toplevel.  Note that (for technical
+instrumentalizing code generation.  See HOL/Decision_Procs/Ferrack for
+a more ambitious application.  In future you ought refrain from ad-hoc
+compiling generated SML code on the ML toplevel.  Note that (for technical
 reasons) @{code} cannot refer to constants for which user-defined
 serializations are set.  Refer to the corresponding ML counterpart
 directly in that cases.
@@ -687,6 +693,12 @@
 Syntax.read_term_global etc.; see also OldGoals.read_term as last
 resort for legacy applications.
 
+* Disposed old declarations, tactics, tactic combinators that refer to
+the simpset or claset of an implicit theory (such as Addsimps,
+Simp_tac, SIMPSET).  INCOMPATIBILITY, should use @{simpset} etc. in
+embedded ML text, or local_simpset_of with a proper context passed as
+explicit runtime argument.
+
 * Antiquotations: block-structured compilation context indicated by
 \<lbrace> ... \<rbrace>; additional antiquotation forms:
 
--- a/doc-src/Classes/Thy/Classes.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/Classes/Thy/Classes.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -458,7 +458,7 @@
   of monoids for lists:
 *}
 
-interpretation %quote list_monoid!: monoid append "[]"
+interpretation %quote list_monoid: monoid append "[]"
   proof qed auto
 
 text {*
@@ -473,7 +473,7 @@
   "replicate 0 _ = []"
   | "replicate (Suc n) xs = xs @ replicate n xs"
 
-interpretation %quote list_monoid!: monoid append "[]" where
+interpretation %quote list_monoid: monoid append "[]" where
   "monoid.pow_nat append [] = replicate"
 proof -
   interpret monoid append "[]" ..
--- a/doc-src/Classes/Thy/document/Classes.tex	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/Classes/Thy/document/Classes.tex	Sat Mar 28 00:13:01 2009 +0100
@@ -863,7 +863,7 @@
 %
 \isatagquote
 \isacommand{interpretation}\isamarkupfalse%
-\ list{\isacharunderscore}monoid{\isacharbang}{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
+\ list{\isacharunderscore}monoid{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{proof}\isamarkupfalse%
 \ \isacommand{qed}\isamarkupfalse%
 \ auto%
@@ -894,7 +894,7 @@
 \ \ {\isacharbar}\ {\isachardoublequoteopen}replicate\ {\isacharparenleft}Suc\ n{\isacharparenright}\ xs\ {\isacharequal}\ xs\ {\isacharat}\ replicate\ n\ xs{\isachardoublequoteclose}\isanewline
 \isanewline
 \isacommand{interpretation}\isamarkupfalse%
-\ list{\isacharunderscore}monoid{\isacharbang}{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+\ list{\isacharunderscore}monoid{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 \ \ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ replicate{\isachardoublequoteclose}\isanewline
 \isacommand{proof}\isamarkupfalse%
 \ {\isacharminus}\isanewline
@@ -1191,7 +1191,7 @@
 \hspace*{0pt}\\
 \hspace*{0pt}pow{\char95}nat ::~forall a.~(Monoid a) => Nat -> a -> a;\\
 \hspace*{0pt}pow{\char95}nat Zero{\char95}nat x = neutral;\\
-\hspace*{0pt}pow{\char95}nat (Suc n) x = mult x (pow{\char95}nat n x);\\
+\hspace*{0pt}pow{\char95}nat (Suc n) xa = mult xa (pow{\char95}nat n xa);\\
 \hspace*{0pt}\\
 \hspace*{0pt}pow{\char95}int ::~forall a.~(Group a) => Integer -> a -> a;\\
 \hspace*{0pt}pow{\char95}int k x =\\
@@ -1272,8 +1272,8 @@
 \hspace*{0pt} ~IntInf.int group;\\
 \hspace*{0pt}\\
 \hspace*{0pt}fun pow{\char95}nat A{\char95}~Zero{\char95}nat x = neutral (monoidl{\char95}monoid A{\char95})\\
-\hspace*{0pt} ~| pow{\char95}nat A{\char95}~(Suc n) x =\\
-\hspace*{0pt} ~~~mult ((semigroup{\char95}monoidl o monoidl{\char95}monoid) A{\char95}) x (pow{\char95}nat A{\char95}~n x);\\
+\hspace*{0pt} ~| pow{\char95}nat A{\char95}~(Suc n) xa =\\
+\hspace*{0pt} ~~~mult ((semigroup{\char95}monoidl o monoidl{\char95}monoid) A{\char95}) xa (pow{\char95}nat A{\char95}~n xa);\\
 \hspace*{0pt}\\
 \hspace*{0pt}fun pow{\char95}int A{\char95}~k x =\\
 \hspace*{0pt} ~(if IntInf.<= ((0 :~IntInf.int),~k)\\
--- a/doc-src/Codegen/Makefile	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/Codegen/Makefile	Sat Mar 28 00:13:01 2009 +0100
@@ -17,7 +17,7 @@
 
 dvi: $(NAME).dvi
 
-$(NAME).dvi: $(FILES) isabelle_isar.eps codegen_process.ps
+$(NAME).dvi: $(FILES) isabelle_isar.eps Thy/pictures/architecture.eps Thy/pictures/adaption.eps
 	$(LATEX) $(NAME)
 	$(BIBTEX) $(NAME)
 	$(LATEX) $(NAME)
@@ -25,7 +25,7 @@
 
 pdf: $(NAME).pdf
 
-$(NAME).pdf: $(FILES) isabelle_isar.pdf codegen_process.pdf
+$(NAME).pdf: $(FILES) isabelle_isar.pdf Thy/pictures/architecture.pdf Thy/pictures/adaption.pdf
 	$(PDFLATEX) $(NAME)
 	$(BIBTEX) $(NAME)
 	$(PDFLATEX) $(NAME)
@@ -33,3 +33,12 @@
 	$(FIXBOOKMARKS) $(NAME).out
 	$(PDFLATEX) $(NAME)
 	$(PDFLATEX) $(NAME)
+
+Thy/pictures/%.dvi: Thy/pictures/%.tex
+	latex -output-directory=$(dir $@) $<
+
+Thy/pictures/%.eps: Thy/pictures/%.dvi
+	dvips -E -o $@ $<
+
+Thy/pictures/%.pdf: Thy/pictures/%.eps
+	epstopdf --outfile=$@ $<
--- a/doc-src/Codegen/Thy/ROOT.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/Codegen/Thy/ROOT.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,3 @@
-
-(* $Id$ *)
 
 no_document use_thy "Setup";
 no_document use_thys ["Efficient_Nat"];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/Thy/pictures/adaption.tex	Sat Mar 28 00:13:01 2009 +0100
@@ -0,0 +1,46 @@
+
+\documentclass[12pt]{article}
+\usepackage{pgf}
+\usepackage{pgflibraryshapes}
+\usepackage{tikz}
+
+\begin{document}
+
+\begin{tikzpicture}[scale = 0.5]
+  \tikzstyle water=[color = blue, thick]
+  \tikzstyle ice=[color = black, very thick, cap = round, join = round, fill = white]
+  \tikzstyle process=[color = green, semithick, ->]
+  \tikzstyle adaption=[color = red, semithick, ->]
+  \tikzstyle target=[color = black]
+  \foreach \x in {0, ..., 24}
+    \draw[style=water] (\x, 0.25) sin + (0.25, 0.25) cos + (0.25, -0.25) sin
+      + (0.25, -0.25) cos + (0.25, 0.25);
+  \draw[style=ice] (1, 0) --
+    (3, 6) node[above, fill=white] {logic} -- (5, 0) -- cycle;
+  \draw[style=ice] (9, 0) --
+    (11, 6) node[above, fill=white] {intermediate language} -- (13, 0) -- cycle;
+  \draw[style=ice] (15, -6) --
+    (19, 6) node[above, fill=white] {target language} -- (23, -6) -- cycle;
+  \draw[style=process]
+    (3.5, 3) .. controls (7, 5) .. node[fill=white] {translation} (10.5, 3);
+  \draw[style=process]
+    (11.5, 3) .. controls (15, 5) .. node[fill=white] (serialisation) {serialisation} (18.5, 3);
+  \node (adaption) at (11, -2) [style=adaption] {adaption};
+  \node at (19, 3) [rotate=90] {generated};
+  \node at (19.5, -5) {language};
+  \node at (19.5, -3) {library};
+  \node (includes) at (19.5, -1) {includes};
+  \node (reserved) at (16.5, -3) [rotate=72] {reserved}; % proper 71.57
+  \draw[style=process]
+    (includes) -- (serialisation);
+  \draw[style=process]
+    (reserved) -- (serialisation);
+  \draw[style=adaption]
+    (adaption) -- (serialisation);
+  \draw[style=adaption]
+    (adaption) -- (includes);
+  \draw[style=adaption]
+    (adaption) -- (reserved);
+\end{tikzpicture}
+
+\end{document}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/Thy/pictures/architecture.tex	Sat Mar 28 00:13:01 2009 +0100
@@ -0,0 +1,33 @@
+
+\documentclass[12pt]{article}
+\usepackage{pgf}
+\usepackage{pgflibraryshapes}
+\usepackage{tikz}
+
+\begin{document}
+
+\begin{tikzpicture}[x = 4.2cm, y = 1cm]
+  \tikzstyle entity=[rounded corners, draw, thick, color = black, fill = white];
+  \tikzstyle process=[ellipse, draw, thick, color = green, fill = white];
+  \tikzstyle process_arrow=[->, semithick, color = green];
+  \node (HOL) at (0, 4) [style=entity] {Isabelle/HOL theory};
+  \node (eqn) at (2, 2) [style=entity] {code equations};
+  \node (iml) at (2, 0) [style=entity] {intermediate language};
+  \node (seri) at (1, 0) [style=process] {serialisation};
+  \node (SML) at (0, 3) [style=entity] {SML};
+  \node (OCaml) at (0, 2) [style=entity] {OCaml};
+  \node (further) at (0, 1) [style=entity] {\ldots};
+  \node (Haskell) at (0, 0) [style=entity] {Haskell};
+  \draw [style=process_arrow] (HOL) .. controls (2, 4) ..
+    node [style=process, near start] {selection}
+    node [style=process, near end] {preprocessing}
+    (eqn);
+  \draw [style=process_arrow] (eqn) -- node (transl) [style=process] {translation} (iml);
+  \draw [style=process_arrow] (iml) -- (seri);
+  \draw [style=process_arrow] (seri) -- (SML);
+  \draw [style=process_arrow] (seri) -- (OCaml);
+  \draw [style=process_arrow, dashed] (seri) -- (further);
+  \draw [style=process_arrow] (seri) -- (Haskell);
+\end{tikzpicture}
+
+\end{document}
Binary file doc-src/Codegen/codegen_process.pdf has changed
--- a/doc-src/Codegen/codegen_process.ps	Sat Mar 28 00:11:02 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,586 +0,0 @@
-%!PS-Adobe-2.0
-%%Creator: dot version 2.2 (Mon Sep 12 23:33:36 UTC 2005)
-%%For: (haftmann) Florian Haftmann
-%%Title: _anonymous_0
-%%Pages: (atend)
-%%BoundingBox: 35 35 451 291
-%%EndComments
-save
-%%BeginProlog
-/DotDict 200 dict def
-DotDict begin
-
-/setupLatin1 {
-mark
-/EncodingVector 256 array def
- EncodingVector 0
-
-ISOLatin1Encoding 0 255 getinterval putinterval
-
-EncodingVector
-  dup 306 /AE
-  dup 301 /Aacute
-  dup 302 /Acircumflex
-  dup 304 /Adieresis
-  dup 300 /Agrave
-  dup 305 /Aring
-  dup 303 /Atilde
-  dup 307 /Ccedilla
-  dup 311 /Eacute
-  dup 312 /Ecircumflex
-  dup 313 /Edieresis
-  dup 310 /Egrave
-  dup 315 /Iacute
-  dup 316 /Icircumflex
-  dup 317 /Idieresis
-  dup 314 /Igrave
-  dup 334 /Udieresis
-  dup 335 /Yacute
-  dup 376 /thorn
-  dup 337 /germandbls
-  dup 341 /aacute
-  dup 342 /acircumflex
-  dup 344 /adieresis
-  dup 346 /ae
-  dup 340 /agrave
-  dup 345 /aring
-  dup 347 /ccedilla
-  dup 351 /eacute
-  dup 352 /ecircumflex
-  dup 353 /edieresis
-  dup 350 /egrave
-  dup 355 /iacute
-  dup 356 /icircumflex
-  dup 357 /idieresis
-  dup 354 /igrave
-  dup 360 /dcroat
-  dup 361 /ntilde
-  dup 363 /oacute
-  dup 364 /ocircumflex
-  dup 366 /odieresis
-  dup 362 /ograve
-  dup 365 /otilde
-  dup 370 /oslash
-  dup 372 /uacute
-  dup 373 /ucircumflex
-  dup 374 /udieresis
-  dup 371 /ugrave
-  dup 375 /yacute
-  dup 377 /ydieresis  
-
-% Set up ISO Latin 1 character encoding
-/starnetISO {
-        dup dup findfont dup length dict begin
-        { 1 index /FID ne { def }{ pop pop } ifelse
-        } forall
-        /Encoding EncodingVector def
-        currentdict end definefont
-} def
-/Times-Roman starnetISO def
-/Times-Italic starnetISO def
-/Times-Bold starnetISO def
-/Times-BoldItalic starnetISO def
-/Helvetica starnetISO def
-/Helvetica-Oblique starnetISO def
-/Helvetica-Bold starnetISO def
-/Helvetica-BoldOblique starnetISO def
-/Courier starnetISO def
-/Courier-Oblique starnetISO def
-/Courier-Bold starnetISO def
-/Courier-BoldOblique starnetISO def
-cleartomark
-} bind def
-
-%%BeginResource: procset graphviz 0 0
-/coord-font-family /Times-Roman def
-/default-font-family /Times-Roman def
-/coordfont coord-font-family findfont 8 scalefont def
-
-/InvScaleFactor 1.0 def
-/set_scale {
-	dup 1 exch div /InvScaleFactor exch def
-	dup scale
-} bind def
-
-% styles
-/solid { [] 0 setdash } bind def
-/dashed { [9 InvScaleFactor mul dup ] 0 setdash } bind def
-/dotted { [1 InvScaleFactor mul 6 InvScaleFactor mul] 0 setdash } bind def
-/invis {/fill {newpath} def /stroke {newpath} def /show {pop newpath} def} bind def
-/bold { 2 setlinewidth } bind def
-/filled { } bind def
-/unfilled { } bind def
-/rounded { } bind def
-/diagonals { } bind def
-
-% hooks for setting color 
-/nodecolor { sethsbcolor } bind def
-/edgecolor { sethsbcolor } bind def
-/graphcolor { sethsbcolor } bind def
-/nopcolor {pop pop pop} bind def
-
-/beginpage {	% i j npages
-	/npages exch def
-	/j exch def
-	/i exch def
-	/str 10 string def
-	npages 1 gt {
-		gsave
-			coordfont setfont
-			0 0 moveto
-			(\() show i str cvs show (,) show j str cvs show (\)) show
-		grestore
-	} if
-} bind def
-
-/set_font {
-	findfont exch
-	scalefont setfont
-} def
-
-% draw aligned label in bounding box aligned to current point
-/alignedtext {			% width adj text
-	/text exch def
-	/adj exch def
-	/width exch def
-	gsave
-		width 0 gt {
-			text stringwidth pop adj mul 0 rmoveto
-		} if
-		[] 0 setdash
-		text show
-	grestore
-} def
-
-/boxprim {				% xcorner ycorner xsize ysize
-		4 2 roll
-		moveto
-		2 copy
-		exch 0 rlineto
-		0 exch rlineto
-		pop neg 0 rlineto
-		closepath
-} bind def
-
-/ellipse_path {
-	/ry exch def
-	/rx exch def
-	/y exch def
-	/x exch def
-	matrix currentmatrix
-	newpath
-	x y translate
-	rx ry scale
-	0 0 1 0 360 arc
-	setmatrix
-} bind def
-
-/endpage { showpage } bind def
-/showpage { } def
-
-/layercolorseq
-	[	% layer color sequence - darkest to lightest
-		[0 0 0]
-		[.2 .8 .8]
-		[.4 .8 .8]
-		[.6 .8 .8]
-		[.8 .8 .8]
-	]
-def
-
-/layerlen layercolorseq length def
-
-/setlayer {/maxlayer exch def /curlayer exch def
-	layercolorseq curlayer 1 sub layerlen mod get
-	aload pop sethsbcolor
-	/nodecolor {nopcolor} def
-	/edgecolor {nopcolor} def
-	/graphcolor {nopcolor} def
-} bind def
-
-/onlayer { curlayer ne {invis} if } def
-
-/onlayers {
-	/myupper exch def
-	/mylower exch def
-	curlayer mylower lt
-	curlayer myupper gt
-	or
-	{invis} if
-} def
-
-/curlayer 0 def
-
-%%EndResource
-%%EndProlog
-%%BeginSetup
-14 default-font-family set_font
-1 setmiterlimit
-% /arrowlength 10 def
-% /arrowwidth 5 def
-
-% make sure pdfmark is harmless for PS-interpreters other than Distiller
-/pdfmark where {pop} {userdict /pdfmark /cleartomark load put} ifelse
-% make '<<' and '>>' safe on PS Level 1 devices
-/languagelevel where {pop languagelevel}{1} ifelse
-2 lt {
-    userdict (<<) cvn ([) cvn load put
-    userdict (>>) cvn ([) cvn load put
-} if
-
-%%EndSetup
-%%Page: 1 1
-%%PageBoundingBox: 36 36 451 291
-%%PageOrientation: Portrait
-gsave
-35 35 416 256 boxprim clip newpath
-36 36 translate
-0 0 1 beginpage
-0 0 translate 0 rotate
-[ /CropBox [36 36 451 291] /PAGES pdfmark
-0.000 0.000 0.000 graphcolor
-14.00 /Times-Roman set_font
-
-%	theory
-gsave 10 dict begin
-newpath 93 254 moveto
-1 254 lineto
-1 214 lineto
-93 214 lineto
-closepath
-stroke
-gsave 10 dict begin
-8 237 moveto
-(Isabelle/HOL)
-[4.56 5.52 6.24 6.96 6.24 3.84 3.84 6.24 3.84 10.08 10.08 8.64]
-xshow
-16 221 moveto
-(Isar theory)
-[4.56 5.52 6.24 4.56 3.6 4.08 6.96 6.24 6.96 4.8 6.96]
-xshow
-end grestore
-end grestore
-
-%	selection
-gsave 10 dict begin
-183 234 38 18 ellipse_path
-stroke
-gsave 10 dict begin
-158 229 moveto
-(selection)
-[5.52 6.24 3.84 6.24 6.24 3.84 3.84 6.96 6.96]
-xshow
-end grestore
-end grestore
-
-%	theory -> selection
-newpath 94 234 moveto
-107 234 121 234 135 234 curveto
-stroke
-gsave 10 dict begin
-solid
-1 setlinewidth
-0.000 0.000 0.000 edgecolor
-newpath 135 238 moveto
-145 234 lineto
-135 231 lineto
-closepath
-fill
-0.000 0.000 0.000 edgecolor
-newpath 135 238 moveto
-145 234 lineto
-135 231 lineto
-closepath
-stroke
-end grestore
-
-%	sml
-gsave 10 dict begin
-newpath 74 144 moveto
-20 144 lineto
-20 108 lineto
-74 108 lineto
-closepath
-stroke
-gsave 10 dict begin
-32 121 moveto
-(SML)
-[7.68 12.48 8.64]
-xshow
-end grestore
-end grestore
-
-%	other
-gsave 10 dict begin
-gsave 10 dict begin
-41 67 moveto
-(...)
-[3.6 3.6 3.6]
-xshow
-end grestore
-end grestore
-
-%	haskell
-gsave 10 dict begin
-newpath 77 36 moveto
-17 36 lineto
-17 0 lineto
-77 0 lineto
-closepath
-stroke
-gsave 10 dict begin
-25 13 moveto
-(Haskell)
-[10.08 6.24 5.52 6.72 6.24 3.84 3.84]
-xshow
-end grestore
-end grestore
-
-%	preprocessing
-gsave 10 dict begin
-183 180 52 18 ellipse_path
-stroke
-gsave 10 dict begin
-143 175 moveto
-(preprocessing)
-[6.96 4.56 6.24 6.96 4.56 6.96 6.24 6.24 5.52 5.52 3.84 6.96 6.96]
-xshow
-end grestore
-end grestore
-
-%	selection -> preprocessing
-newpath 183 216 moveto
-183 213 183 211 183 208 curveto
-stroke
-gsave 10 dict begin
-solid
-1 setlinewidth
-0.000 0.000 0.000 edgecolor
-newpath 187 208 moveto
-183 198 lineto
-180 208 lineto
-closepath
-fill
-0.000 0.000 0.000 edgecolor
-newpath 187 208 moveto
-183 198 lineto
-180 208 lineto
-closepath
-stroke
-end grestore
-
-%	def_eqn
-gsave 10 dict begin
-newpath 403 198 moveto
-283 198 lineto
-283 162 lineto
-403 162 lineto
-closepath
-stroke
-gsave 10 dict begin
-291 175 moveto
-(defining equations)
-[6.96 6.24 4.8 3.84 6.96 3.84 6.96 6.96 3.6 6.24 6.72 6.96 6.24 3.84 3.84 6.96 6.96 5.52]
-xshow
-end grestore
-end grestore
-
-%	preprocessing -> def_eqn
-newpath 236 180 moveto
-248 180 260 180 273 180 curveto
-stroke
-gsave 10 dict begin
-solid
-1 setlinewidth
-0.000 0.000 0.000 edgecolor
-newpath 273 184 moveto
-283 180 lineto
-273 177 lineto
-closepath
-fill
-0.000 0.000 0.000 edgecolor
-newpath 273 184 moveto
-283 180 lineto
-273 177 lineto
-closepath
-stroke
-end grestore
-
-%	serialization
-gsave 10 dict begin
-183 72 47 18 ellipse_path
-stroke
-gsave 10 dict begin
-148 67 moveto
-(serialization)
-[5.52 6.24 4.8 3.84 6.24 3.84 3.84 6.24 6.24 3.84 3.84 6.96 6.96]
-xshow
-end grestore
-end grestore
-
-%	serialization -> sml
-newpath 150 85 moveto
-129 93 104 103 83 111 curveto
-stroke
-gsave 10 dict begin
-solid
-1 setlinewidth
-0.000 0.000 0.000 edgecolor
-newpath 82 108 moveto
-74 115 lineto
-85 114 lineto
-closepath
-fill
-0.000 0.000 0.000 edgecolor
-newpath 82 108 moveto
-74 115 lineto
-85 114 lineto
-closepath
-stroke
-end grestore
-
-%	serialization -> other
-gsave 10 dict begin
-dotted
-newpath 135 72 moveto
-119 72 100 72 84 72 curveto
-stroke
-gsave 10 dict begin
-solid
-1 setlinewidth
-0.000 0.000 0.000 edgecolor
-newpath 84 69 moveto
-74 72 lineto
-84 76 lineto
-closepath
-fill
-0.000 0.000 0.000 edgecolor
-newpath 84 69 moveto
-74 72 lineto
-84 76 lineto
-closepath
-stroke
-end grestore
-end grestore
-
-%	serialization -> haskell
-newpath 150 59 moveto
-131 51 107 42 86 34 curveto
-stroke
-gsave 10 dict begin
-solid
-1 setlinewidth
-0.000 0.000 0.000 edgecolor
-newpath 88 31 moveto
-77 30 lineto
-85 37 lineto
-closepath
-fill
-0.000 0.000 0.000 edgecolor
-newpath 88 31 moveto
-77 30 lineto
-85 37 lineto
-closepath
-stroke
-end grestore
-
-%	translation
-gsave 10 dict begin
-343 126 43 18 ellipse_path
-stroke
-gsave 10 dict begin
-313 121 moveto
-(translation)
-[3.84 4.56 6.24 6.96 5.52 3.84 6.24 3.84 3.84 6.96 6.96]
-xshow
-end grestore
-end grestore
-
-%	def_eqn -> translation
-newpath 343 162 moveto
-343 159 343 157 343 154 curveto
-stroke
-gsave 10 dict begin
-solid
-1 setlinewidth
-0.000 0.000 0.000 edgecolor
-newpath 347 154 moveto
-343 144 lineto
-340 154 lineto
-closepath
-fill
-0.000 0.000 0.000 edgecolor
-newpath 347 154 moveto
-343 144 lineto
-340 154 lineto
-closepath
-stroke
-end grestore
-
-%	iml
-gsave 10 dict begin
-newpath 413 90 moveto
-273 90 lineto
-273 54 lineto
-413 54 lineto
-closepath
-stroke
-gsave 10 dict begin
-280 67 moveto
-(intermediate language)
-[3.84 6.96 3.84 6.24 4.8 10.8 6.24 6.96 3.84 6.24 3.84 6.24 3.6 3.84 6.24 6.96 6.96 6.96 6.24 6.72 6.24]
-xshow
-end grestore
-end grestore
-
-%	translation -> iml
-newpath 343 108 moveto
-343 105 343 103 343 100 curveto
-stroke
-gsave 10 dict begin
-solid
-1 setlinewidth
-0.000 0.000 0.000 edgecolor
-newpath 347 100 moveto
-343 90 lineto
-340 100 lineto
-closepath
-fill
-0.000 0.000 0.000 edgecolor
-newpath 347 100 moveto
-343 90 lineto
-340 100 lineto
-closepath
-stroke
-end grestore
-
-%	iml -> serialization
-newpath 272 72 moveto
-262 72 251 72 241 72 curveto
-stroke
-gsave 10 dict begin
-solid
-1 setlinewidth
-0.000 0.000 0.000 edgecolor
-newpath 241 69 moveto
-231 72 lineto
-241 76 lineto
-closepath
-fill
-0.000 0.000 0.000 edgecolor
-newpath 241 69 moveto
-231 72 lineto
-241 76 lineto
-closepath
-stroke
-end grestore
-endpage
-showpage
-grestore
-%%PageTrailer
-%%EndPage: 1
-%%Trailer
-%%Pages: 1
-end
-restore
-%%EOF
--- a/doc-src/HOL/HOL.tex	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/HOL/HOL.tex	Sat Mar 28 00:13:01 2009 +0100
@@ -1427,7 +1427,7 @@
 provides a decision procedure for \emph{linear arithmetic}: formulae involving
 addition and subtraction. The simplifier invokes a weak version of this
 decision procedure automatically. If this is not sufficent, you can invoke the
-full procedure \ttindex{arith_tac} explicitly.  It copes with arbitrary
+full procedure \ttindex{linear_arith_tac} explicitly.  It copes with arbitrary
 formulae involving {\tt=}, {\tt<}, {\tt<=}, {\tt+}, {\tt-}, {\tt Suc}, {\tt
   min}, {\tt max} and numerical constants. Other subterms are treated as
 atomic, while subformulae not involving numerical types are ignored. Quantified
@@ -1438,10 +1438,10 @@
 If {\tt k} is a numeral, then {\tt div k}, {\tt mod k} and
 {\tt k dvd} are also supported. The former two are eliminated
 by case distinctions, again blowing up the running time.
-If the formula involves explicit quantifiers, \texttt{arith_tac} may take
+If the formula involves explicit quantifiers, \texttt{linear_arith_tac} may take
 super-exponential time and space.
 
-If \texttt{arith_tac} fails, try to find relevant arithmetic results in
+If \texttt{linear_arith_tac} fails, try to find relevant arithmetic results in
 the library.  The theories \texttt{Nat} and \texttt{NatArith} contain
 theorems about {\tt<}, {\tt<=}, \texttt{+}, \texttt{-} and \texttt{*}.
 Theory \texttt{Divides} contains theorems about \texttt{div} and
--- a/doc-src/IsarOverview/Isar/Induction.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/IsarOverview/Isar/Induction.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -143,14 +143,14 @@
 universally quantifies all \emph{vars} before the induction.  Hence
 they can be replaced by \emph{arbitrary} values in the proof.
 
-The nice thing about generalization via @{text"arbitrary:"} is that in
-the induction step the claim is available in unquantified form but
+Generalization via @{text"arbitrary"} is particularly convenient
+if the induction step is a structured proof as opposed to the automatic
+example above. Then the claim is available in unquantified form but
 with the generalized variables replaced by @{text"?"}-variables, ready
-for instantiation. In the above example the
-induction hypothesis is @{text"itrev xs ?ys = rev xs @ ?ys"}.
+for instantiation. In the above example, in the @{const[source] Cons} case the
+induction hypothesis is @{text"itrev xs ?ys = rev xs @ ?ys"} (available
+under the name @{const[source] Cons}).
 
-For the curious: @{text"arbitrary:"} introduces @{text"\<And>"}
-behind the scenes.
 
 \subsection{Inductive proofs of conditional formulae}
 \label{sec:full-Ind}
--- a/doc-src/IsarOverview/Isar/Logic.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/IsarOverview/Isar/Logic.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -30,8 +30,8 @@
   show A by(rule a)
 qed
 
-text{*\noindent Single-identifier formulae such as @{term A} need not
-be enclosed in double quotes. However, we will continue to do so for
+text{*\noindent As you see above, single-identifier formulae such as @{term A}
+need not be enclosed in double quotes. However, we will continue to do so for
 uniformity.
 
 Instead of applying fact @{text a} via the @{text rule} method, we can
--- a/doc-src/IsarOverview/Isar/document/Induction.tex	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/IsarOverview/Isar/document/Induction.tex	Sat Mar 28 00:13:01 2009 +0100
@@ -342,14 +342,14 @@
 universally quantifies all \emph{vars} before the induction.  Hence
 they can be replaced by \emph{arbitrary} values in the proof.
 
-The nice thing about generalization via \isa{arbitrary{\isacharcolon}} is that in
-the induction step the claim is available in unquantified form but
+Generalization via \isa{arbitrary} is particularly convenient
+if the induction step is a structured proof as opposed to the automatic
+example above. Then the claim is available in unquantified form but
 with the generalized variables replaced by \isa{{\isacharquery}}-variables, ready
-for instantiation. In the above example the
-induction hypothesis is \isa{itrev\ xs\ {\isacharquery}ys\ {\isacharequal}\ rev\ xs\ {\isacharat}\ {\isacharquery}ys}.
+for instantiation. In the above example, in the \isa{Cons} case the
+induction hypothesis is \isa{itrev\ xs\ {\isacharquery}ys\ {\isacharequal}\ rev\ xs\ {\isacharat}\ {\isacharquery}ys} (available
+under the name \isa{Cons}).
 
-For the curious: \isa{arbitrary{\isacharcolon}} introduces \isa{{\isasymAnd}}
-behind the scenes.
 
 \subsection{Inductive proofs of conditional formulae}
 \label{sec:full-Ind}
--- a/doc-src/IsarOverview/Isar/document/Logic.tex	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/IsarOverview/Isar/document/Logic.tex	Sat Mar 28 00:13:01 2009 +0100
@@ -93,8 +93,8 @@
 \endisadelimproof
 %
 \begin{isamarkuptext}%
-\noindent Single-identifier formulae such as \isa{A} need not
-be enclosed in double quotes. However, we will continue to do so for
+\noindent As you see above, single-identifier formulae such as \isa{A}
+need not be enclosed in double quotes. However, we will continue to do so for
 uniformity.
 
 Instead of applying fact \isa{a} via the \isa{rule} method, we can
--- a/doc-src/TutorialI/Documents/Documents.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/TutorialI/Documents/Documents.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -617,7 +617,7 @@
   same types as they have in the main goal statement.
 
   \medskip Several further kinds of antiquotations and options are
-  available \cite{isabelle-sys}.  Here are a few commonly used
+  available \cite{isabelle-isar-ref}.  Here are a few commonly used
   combinations:
 
   \medskip
--- a/doc-src/TutorialI/Documents/document/Documents.tex	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/TutorialI/Documents/document/Documents.tex	Sat Mar 28 00:13:01 2009 +0100
@@ -691,7 +691,7 @@
   same types as they have in the main goal statement.
 
   \medskip Several further kinds of antiquotations and options are
-  available \cite{isabelle-sys}.  Here are a few commonly used
+  available \cite{isabelle-isar-ref}.  Here are a few commonly used
   combinations:
 
   \medskip
--- a/doc-src/TutorialI/Protocol/Message.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/TutorialI/Protocol/Message.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Auth/Message
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1996  University of Cambridge
 
@@ -830,9 +829,9 @@
 (*Prove base case (subgoal i) and simplify others.  A typical base case
   concerns  Crypt K X \<notin> Key`shrK`bad  and cannot be proved by rewriting
   alone.*)
-fun prove_simple_subgoals_tac i = 
-    force_tac (claset(), simpset() addsimps [@{thm image_eq_UN}]) i THEN
-    ALLGOALS Asm_simp_tac
+fun prove_simple_subgoals_tac (cs, ss) i = 
+    force_tac (cs, ss addsimps [@{thm image_eq_UN}]) i THEN
+    ALLGOALS (asm_simp_tac ss)
 
 (*Analysis of Fake cases.  Also works for messages that forward unknown parts,
   but this application is no longer necessary if analz_insert_eq is used.
@@ -857,8 +856,7 @@
 		  (cs addIs [analz_insertI,
 				   impOfSubs analz_subset_parts]) 4 1))
 
-(*The explicit claset and simpset arguments help it work with Isar*)
-fun gen_spy_analz_tac (cs,ss) i =
+fun spy_analz_tac (cs,ss) i =
   DETERM
    (SELECT_GOAL
      (EVERY 
@@ -870,8 +868,6 @@
        simp_tac ss 1,
        REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])),
        DEPTH_SOLVE (atomic_spy_analz_tac (cs,ss) 1)]) i)
-
-fun spy_analz_tac i = gen_spy_analz_tac (claset(), simpset()) i
 *}
 
 text{*By default only @{text o_apply} is built-in.  But in the presence of
@@ -919,7 +915,7 @@
 lemmas Fake_parts_sing_imp_Un = Fake_parts_sing [THEN [2] rev_subsetD]
 
 method_setup spy_analz = {*
-    Scan.succeed (SIMPLE_METHOD' o gen_spy_analz_tac o local_clasimpset_of) *}
+    Scan.succeed (SIMPLE_METHOD' o spy_analz_tac o local_clasimpset_of) *}
     "for proving the Fake case when analz is involved"
 
 method_setup atomic_spy_analz = {*
--- a/doc-src/TutorialI/Protocol/Public.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/TutorialI/Protocol/Public.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Auth/Public
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1996  University of Cambridge
 
@@ -153,15 +152,15 @@
 
 (*Tactic for possibility theorems*)
 ML {*
-fun possibility_tac st = st |>
+fun possibility_tac ctxt =
     REPEAT (*omit used_Says so that Nonces start from different traces!*)
-    (ALLGOALS (simp_tac (@{simpset} delsimps [used_Says]))
+    (ALLGOALS (simp_tac (local_simpset_of ctxt delsimps [used_Says]))
      THEN
      REPEAT_FIRST (eq_assume_tac ORELSE' 
                    resolve_tac [refl, conjI, @{thm Nonce_supply}]));
 *}
 
-method_setup possibility = {* Scan.succeed (K (SIMPLE_METHOD possibility_tac)) *}
+method_setup possibility = {* Scan.succeed (SIMPLE_METHOD o possibility_tac) *}
     "for proving possibility theorems"
 
 end
--- a/doc-src/TutorialI/Rules/rules.tex	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/TutorialI/Rules/rules.tex	Sat Mar 28 00:13:01 2009 +0100
@@ -2138,11 +2138,11 @@
 
 \index{*insert (method)|(}%
 The \isa{insert} method
-inserts a given theorem as a new assumption of the current subgoal.  This
+inserts a given theorem as a new assumption of all subgoals.  This
 already is a forward step; moreover, we may (as always when using a
 theorem) apply
 \isa{of}, \isa{THEN} and other directives.  The new assumption can then
-be used to help prove the subgoal.
+be used to help prove the subgoals.
 
 For example, consider this theorem about the divides relation.  The first
 proof step inserts the distributive law for
--- a/doc-src/TutorialI/Sets/sets.tex	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/TutorialI/Sets/sets.tex	Sat Mar 28 00:13:01 2009 +0100
@@ -299,7 +299,7 @@
 \isa{UN x:A.\ B} in \textsc{ascii}. Indexed union satisfies this basic law:
 \begin{isabelle}
 (b\ \isasymin\
-(\isasymUnion x\isasymin A. B\ x) =\ (\isasymexists x\isasymin A.\
+(\isasymUnion x\isasymin A. B\ x)) =\ (\isasymexists x\isasymin A.\
 b\ \isasymin\ B\ x)
 \rulenamedx{UN_iff}
 \end{isabelle}
@@ -414,12 +414,12 @@
 $k$-element subsets of~$A$ is \index{binomial coefficients}
 $\binom{n}{k}$.
 
-\begin{warn}
-The term \isa{finite\ A} is defined via a syntax translation as an
-abbreviation for \isa{A {\isasymin} Finites}, where the constant
-\cdx{Finites} denotes the set of all finite sets of a given type.  There
-is no constant \isa{finite}.
-\end{warn}
+%\begin{warn}
+%The term \isa{finite\ A} is defined via a syntax translation as an
+%abbreviation for \isa{A {\isasymin} Finites}, where the constant
+%\cdx{Finites} denotes the set of all finite sets of a given type.  There
+%is no constant \isa{finite}.
+%\end{warn}
 \index{sets|)}
 
 
--- a/doc-src/TutorialI/Types/Overloading2.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/TutorialI/Types/Overloading2.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -15,7 +15,7 @@
               size xs = size ys \<and> (\<forall>i<size xs. xs!i <<= ys!i)"
 
 text{*\noindent
-The infix function @{text"!"} yields the nth element of a list.
+The infix function @{text"!"} yields the nth element of a list, starting with 0.
 
 \begin{warn}
 A type constructor can be instantiated in only one way to
--- a/doc-src/TutorialI/Types/document/Overloading2.tex	Sat Mar 28 00:11:02 2009 +0100
+++ b/doc-src/TutorialI/Types/document/Overloading2.tex	Sat Mar 28 00:13:01 2009 +0100
@@ -46,7 +46,7 @@
 \ \ \ \ \ \ \ \ \ \ \ \ \ \ size\ xs\ {\isacharequal}\ size\ ys\ {\isasymand}\ {\isacharparenleft}{\isasymforall}i{\isacharless}size\ xs{\isachardot}\ xs{\isacharbang}i\ {\isacharless}{\isacharless}{\isacharequal}\ ys{\isacharbang}i{\isacharparenright}{\isachardoublequoteclose}%
 \begin{isamarkuptext}%
 \noindent
-The infix function \isa{{\isacharbang}} yields the nth element of a list.
+The infix function \isa{{\isacharbang}} yields the nth element of a list, starting with 0.
 
 \begin{warn}
 A type constructor can be instantiated in only one way to
--- a/src/CCL/Type.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/CCL/Type.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      CCL/Type.thy
-    ID:         $Id$
     Author:     Martin Coen
     Copyright   1993  University of Cambridge
 *)
@@ -409,7 +408,7 @@
 
 fun mk_genIs thy defs genXH gen_mono s = prove_goalw thy defs s
       (fn prems => [rtac (genXH RS iffD2) 1,
-                    SIMPSET' simp_tac 1,
+                    simp_tac (simpset_of thy) 1,
                     TRY (fast_tac (@{claset} addIs
                             ([genXH RS iffD2,gen_mono RS coinduct3_mono_lemma RS lfpI]
                              @ prems)) 1)])
@@ -442,8 +441,8 @@
    "<[],[]> : POgen(lfp(%x. POgen(x) Un R Un PO))",
    "[| <h,h'> : lfp(%x. POgen(x) Un R Un PO);  <t,t'> : lfp(%x. POgen(x) Un R Un PO) |] ==> <h$t,h'$t'> : POgen(lfp(%x. POgen(x) Un R Un PO))"];
 
-fun POgen_tac (rla,rlb) i =
-  SELECT_GOAL (CLASET safe_tac) i THEN
+fun POgen_tac ctxt (rla,rlb) i =
+  SELECT_GOAL (safe_tac (local_claset_of ctxt)) i THEN
   rtac (rlb RS (rla RS (thm "ssubst_pair"))) i THEN
   (REPEAT (resolve_tac (POgenIs @ [thm "PO_refl" RS (thm "POgen_mono" RS ci3_AI)] @
     (POgenIs RL [thm "POgen_mono" RS ci3_AgenI]) @ [thm "POgen_mono" RS ci3_RI]) i));
--- a/src/FOL/blastdata.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/FOL/blastdata.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,5 @@
 
-(*** Applying BlastFun to create Blast_tac ***)
+(*** Applying BlastFun to create blast_tac ***)
 structure Blast_Data = 
   struct
   type claset	= Cla.claset
@@ -10,13 +10,10 @@
   val contr_tac = Cla.contr_tac
   val dup_intr	= Cla.dup_intr
   val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac
-  val claset	= Cla.claset
-  val rep_cs    = Cla.rep_cs
+  val rep_cs = Cla.rep_cs
   val cla_modifiers = Cla.cla_modifiers;
   val cla_meth' = Cla.cla_meth'
   end;
 
 structure Blast = BlastFun(Blast_Data);
-
-val Blast_tac = Blast.Blast_tac
-and blast_tac = Blast.blast_tac;
+val blast_tac = Blast.blast_tac;
--- a/src/FOL/ex/Iff_Oracle.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/FOL/ex/Iff_Oracle.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -34,7 +34,7 @@
 
 ML {* iff_oracle (@{theory}, 2) *}
 ML {* iff_oracle (@{theory}, 10) *}
-ML {* Thm.proof_of (iff_oracle (@{theory}, 10)) *}
+ML {* Thm.proof_body_of (iff_oracle (@{theory}, 10)) *}
 
 text {* These oracle calls had better fail. *}
 
--- a/src/FOL/ex/LocaleTest.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/FOL/ex/LocaleTest.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -119,7 +119,7 @@
 
 term extra_type.test thm extra_type.test_def
 
-interpretation var: extra_type "0" "%x y. x = 0" .
+interpretation var?: extra_type "0" "%x y. x = 0" .
 
 thm var.test_def
 
@@ -381,13 +381,13 @@
 
 subsection {* Sublocale, then interpretation in theory *}
 
-interpretation int: lgrp "op +" "0" "minus"
+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 +"
+interpretation int2?: semi "op +"
   by unfold_locales  (* subsumed, thm int2.assoc not generated *)
 
 thm int.lone int.right.rone
@@ -443,7 +443,7 @@
 
 end
 
-interpretation x!: logic_o "op &" "Not"
+interpretation x: logic_o "op &" "Not"
   where bool_logic_o: "logic_o.lor_o(op &, Not, x, y) <-> x | y"
 proof -
   show bool_logic_o: "PROP logic_o(op &, Not)" by unfold_locales fast+
--- a/src/FOL/simpdata.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/FOL/simpdata.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -117,8 +117,6 @@
 val split_asm_tac    = Splitter.split_asm_tac;
 val op addsplits     = Splitter.addsplits;
 val op delsplits     = Splitter.delsplits;
-val Addsplits        = Splitter.Addsplits;
-val Delsplits        = Splitter.Delsplits;
 
 
 (*** Standard simpsets ***)
--- a/src/HOL/Algebra/AbelCoset.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Algebra/AbelCoset.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -540,8 +540,8 @@
                                   (| carrier = carrier H, mult = add H, one = zero H |) h"
   shows "abelian_group_hom G H h"
 proof -
-  interpret G!: abelian_group G by fact
-  interpret H!: abelian_group H by fact
+  interpret G: abelian_group G by fact
+  interpret H: abelian_group H by fact
   show ?thesis apply (intro abelian_group_hom.intro abelian_group_hom_axioms.intro)
     apply fact
     apply fact
@@ -692,7 +692,7 @@
   assumes carr: "x \<in> carrier G" "x' \<in> carrier G"
   shows "(x' \<in> H +> x) = (x' \<ominus> x \<in> H)"
 proof -
-  interpret G!: ring G by fact
+  interpret G: ring G by fact
   from carr
   have "(x' \<in> H +> x) = (x' \<oplus> \<ominus>x \<in> H)" by (rule a_rcos_module)
   with carr
--- a/src/HOL/Algebra/Group.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Algebra/Group.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -488,8 +488,8 @@
   assumes "monoid G" and "monoid H"
   shows "monoid (G \<times>\<times> H)"
 proof -
-  interpret G!: monoid G by fact
-  interpret H!: monoid H by fact
+  interpret G: monoid G by fact
+  interpret H: monoid H by fact
   from assms
   show ?thesis by (unfold monoid_def DirProd_def, auto) 
 qed
@@ -500,8 +500,8 @@
   assumes "group G" and "group H"
   shows "group (G \<times>\<times> H)"
 proof -
-  interpret G!: group G by fact
-  interpret H!: group H by fact
+  interpret G: group G by fact
+  interpret H: group H by fact
   show ?thesis by (rule groupI)
      (auto intro: G.m_assoc H.m_assoc G.l_inv H.l_inv
            simp add: DirProd_def)
@@ -525,9 +525,9 @@
       and h: "h \<in> carrier H"
   shows "m_inv (G \<times>\<times> H) (g, h) = (inv\<^bsub>G\<^esub> g, inv\<^bsub>H\<^esub> h)"
 proof -
-  interpret G!: group G by fact
-  interpret H!: group H by fact
-  interpret Prod!: group "G \<times>\<times> H"
+  interpret G: group G by fact
+  interpret H: group H by fact
+  interpret Prod: group "G \<times>\<times> H"
     by (auto intro: DirProd_group group.intro group.axioms assms)
   show ?thesis by (simp add: Prod.inv_equality g h)
 qed
--- a/src/HOL/Algebra/Ideal.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Algebra/Ideal.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -711,7 +711,7 @@
   obtains "carrier R = I"
     | "\<exists>a b. a \<in> carrier R \<and> b \<in> carrier R \<and> a \<otimes> b \<in> I \<and> a \<notin> I \<and> b \<notin> I"
 proof -
-  interpret R!: cring R by fact
+  interpret R: cring R by fact
   assume "carrier R = I ==> thesis"
     and "\<exists>a b. a \<in> carrier R \<and> b \<in> carrier R \<and> a \<otimes> b \<in> I \<and> a \<notin> I \<and> b \<notin> I \<Longrightarrow> thesis"
   then show thesis using primeidealCD [OF R.is_cring notprime] by blast
--- a/src/HOL/Algebra/IntRing.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Algebra/IntRing.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -96,7 +96,7 @@
   interpretation needs to be done as early as possible --- that is,
   with as few assumptions as possible. *}
 
-interpretation int!: monoid \<Z>
+interpretation int: monoid \<Z>
   where "carrier \<Z> = UNIV"
     and "mult \<Z> x y = x * y"
     and "one \<Z> = 1"
@@ -104,7 +104,7 @@
 proof -
   -- "Specification"
   show "monoid \<Z>" proof qed (auto simp: int_ring_def)
-  then interpret int!: monoid \<Z> .
+  then interpret int: monoid \<Z> .
 
   -- "Carrier"
   show "carrier \<Z> = UNIV" by (simp add: int_ring_def)
@@ -116,12 +116,12 @@
   show "pow \<Z> x n = x^n" by (induct n) (simp, simp add: int_ring_def)+
 qed
 
-interpretation int!: comm_monoid \<Z>
+interpretation int: comm_monoid \<Z>
   where "finprod \<Z> f A = (if finite A then setprod f A else undefined)"
 proof -
   -- "Specification"
   show "comm_monoid \<Z>" proof qed (auto simp: int_ring_def)
-  then interpret int!: comm_monoid \<Z> .
+  then interpret int: comm_monoid \<Z> .
 
   -- "Operations"
   { fix x y have "mult \<Z> x y = x * y" by (simp add: int_ring_def) }
@@ -139,14 +139,14 @@
   qed
 qed
 
-interpretation int!: abelian_monoid \<Z>
+interpretation int: abelian_monoid \<Z>
   where "zero \<Z> = 0"
     and "add \<Z> x y = x + y"
     and "finsum \<Z> f A = (if finite A then setsum f A else undefined)"
 proof -
   -- "Specification"
   show "abelian_monoid \<Z>" proof qed (auto simp: int_ring_def)
-  then interpret int!: abelian_monoid \<Z> .
+  then interpret int: abelian_monoid \<Z> .
 
   -- "Operations"
   { fix x y show "add \<Z> x y = x + y" by (simp add: int_ring_def) }
@@ -164,7 +164,7 @@
   qed
 qed
 
-interpretation int!: abelian_group \<Z>
+interpretation int: abelian_group \<Z>
   where "a_inv \<Z> x = - x"
     and "a_minus \<Z> x y = x - y"
 proof -
@@ -174,7 +174,7 @@
     show "!!x. x \<in> carrier \<Z> ==> EX y : carrier \<Z>. y \<oplus>\<^bsub>\<Z>\<^esub> x = \<zero>\<^bsub>\<Z>\<^esub>"
       by (simp add: int_ring_def) arith
   qed (auto simp: int_ring_def)
-  then interpret int!: abelian_group \<Z> .
+  then interpret int: abelian_group \<Z> .
 
   -- "Operations"
   { fix x y have "add \<Z> x y = x + y" by (simp add: int_ring_def) }
@@ -187,7 +187,7 @@
   show "a_minus \<Z> x y = x - y" by (simp add: int.minus_eq add a_inv)
 qed
 
-interpretation int!: "domain" \<Z>
+interpretation int: "domain" \<Z>
   proof qed (auto simp: int_ring_def left_distrib right_distrib)
 
 
@@ -203,7 +203,7 @@
   "(True ==> PROP R) == PROP R"
   by simp_all
 
-interpretation int! (* FIXME [unfolded UNIV] *) :
+interpretation int (* FIXME [unfolded UNIV] *) :
   partial_order "(| carrier = UNIV::int set, eq = op =, le = op \<le> |)"
   where "carrier (| carrier = UNIV::int set, eq = op =, le = op \<le> |) = UNIV"
     and "le (| carrier = UNIV::int set, eq = op =, le = op \<le> |) x y = (x \<le> y)"
@@ -219,7 +219,7 @@
     by (simp add: lless_def) auto
 qed
 
-interpretation int! (* FIXME [unfolded UNIV] *) :
+interpretation int (* FIXME [unfolded UNIV] *) :
   lattice "(| carrier = UNIV::int set, eq = op =, le = op \<le> |)"
   where "join (| carrier = UNIV::int set, eq = op =, le = op \<le> |) x y = max x y"
     and "meet (| carrier = UNIV::int set, eq = op =, le = op \<le> |) x y = min x y"
@@ -232,7 +232,7 @@
     apply (simp add: greatest_def Lower_def)
     apply arith
     done
-  then interpret int!: lattice "?Z" .
+  then interpret int: lattice "?Z" .
   show "join ?Z x y = max x y"
     apply (rule int.joinI)
     apply (simp_all add: least_def Upper_def)
@@ -245,7 +245,7 @@
     done
 qed
 
-interpretation int! (* [unfolded UNIV] *) :
+interpretation int (* [unfolded UNIV] *) :
   total_order "(| carrier = UNIV::int set, eq = op =, le = op \<le> |)"
   proof qed clarsimp
 
--- a/src/HOL/Algebra/RingHom.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Algebra/RingHom.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -61,8 +61,8 @@
   assumes h: "h \<in> ring_hom R S"
   shows "ring_hom_ring R S h"
 proof -
-  interpret R!: ring R by fact
-  interpret S!: ring S by fact
+  interpret R: ring R by fact
+  interpret S: ring S by fact
   show ?thesis apply (intro ring_hom_ring.intro ring_hom_ring_axioms.intro)
     apply (rule R.is_ring)
     apply (rule S.is_ring)
@@ -78,8 +78,8 @@
   shows "ring_hom_ring R S h"
 proof -
   interpret abelian_group_hom R S h by fact
-  interpret R!: ring R by fact
-  interpret S!: ring S by fact
+  interpret R: ring R by fact
+  interpret S: ring S by fact
   show ?thesis apply (intro ring_hom_ring.intro ring_hom_ring_axioms.intro, rule R.is_ring, rule S.is_ring)
     apply (insert group_hom.homh[OF a_group_hom])
     apply (unfold hom_def ring_hom_def, simp)
@@ -94,8 +94,8 @@
   shows "ring_hom_cring R S h"
 proof -
   interpret ring_hom_ring R S h by fact
-  interpret R!: cring R by fact
-  interpret S!: cring S by fact
+  interpret R: cring R by fact
+  interpret S: cring S by fact
   show ?thesis by (intro ring_hom_cring.intro ring_hom_cring_axioms.intro)
     (rule R.is_cring, rule S.is_cring, rule homh)
 qed
--- a/src/HOL/Algebra/UnivPoly.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Algebra/UnivPoly.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1886,7 +1886,7 @@
   "UP INTEG"} globally.
 *}
 
-interpretation INTEG!: UP_pre_univ_prop INTEG INTEG id "UP INTEG"
+interpretation INTEG: UP_pre_univ_prop INTEG INTEG id "UP INTEG"
   using INTEG_id_eval by simp_all
 
 lemma INTEG_closed [intro, simp]:
--- a/src/HOL/Algebra/ringsimp.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Algebra/ringsimp.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -62,11 +62,13 @@
   Thm.declaration_attribute
     (fn _ => Data.map (AList.delete struct_eq s));
 
-val attribute =
-  Scan.lift ((Args.add >> K true || Args.del >> K false) --| Args.colon ||
-    Scan.succeed true) -- Scan.lift Args.name --
-  Scan.repeat Args.term
-  >> (fn ((b, n), ts) => if b then add_struct_thm (n, ts) else del_struct (n, ts));
+val attrib_setup =
+  Attrib.setup @{binding algebra}
+    (Scan.lift ((Args.add >> K true || Args.del >> K false) --| Args.colon || Scan.succeed true)
+      -- Scan.lift Args.name -- Scan.repeat Args.term
+      >> (fn ((b, n), ts) => if b then add_struct_thm (n, ts) else del_struct (n, ts)))
+    "theorems controlling algebra method";
+
 
 
 (** Setup **)
@@ -74,6 +76,6 @@
 val setup =
   Method.setup @{binding algebra} (Scan.succeed (SIMPLE_METHOD' o algebra_tac))
     "normalisation of algebraic structure" #>
-  Attrib.setup @{binding algebra} attribute "theorems controlling algebra method";
+  attrib_setup;
 
 end;
--- a/src/HOL/Arith_Tools.thy	Sat Mar 28 00:11:02 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,388 +0,0 @@
-(*  Title:      HOL/Arith_Tools.thy
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Author:     Amine Chaieb, TU Muenchen
-*)
-
-header {* Setup of arithmetic tools *}
-
-theory Arith_Tools
-imports NatBin
-uses
-  "~~/src/Provers/Arith/cancel_numeral_factor.ML"
-  "~~/src/Provers/Arith/extract_common_term.ML"
-  "Tools/int_factor_simprocs.ML"
-  "Tools/nat_simprocs.ML"
-  "Tools/Qelim/qelim.ML"
-begin
-
-subsection {* Simprocs for the Naturals *}
-
-declaration {* K nat_simprocs_setup *}
-
-subsubsection{*For simplifying @{term "Suc m - K"} and  @{term "K - Suc m"}*}
-
-text{*Where K above is a literal*}
-
-lemma Suc_diff_eq_diff_pred: "Numeral0 < n ==> Suc m - n = m - (n - Numeral1)"
-by (simp add: numeral_0_eq_0 numeral_1_eq_1 split add: nat_diff_split)
-
-text {*Now just instantiating @{text n} to @{text "number_of v"} does
-  the right simplification, but with some redundant inequality
-  tests.*}
-lemma neg_number_of_pred_iff_0:
-  "neg (number_of (Int.pred v)::int) = (number_of v = (0::nat))"
-apply (subgoal_tac "neg (number_of (Int.pred v)) = (number_of v < Suc 0) ")
-apply (simp only: less_Suc_eq_le le_0_eq)
-apply (subst less_number_of_Suc, simp)
-done
-
-text{*No longer required as a simprule because of the @{text inverse_fold}
-   simproc*}
-lemma Suc_diff_number_of:
-     "Int.Pls < v ==>
-      Suc m - (number_of v) = m - (number_of (Int.pred v))"
-apply (subst Suc_diff_eq_diff_pred)
-apply simp
-apply (simp del: nat_numeral_1_eq_1)
-apply (auto simp only: diff_nat_number_of less_0_number_of [symmetric]
-                        neg_number_of_pred_iff_0)
-done
-
-lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n"
-by (simp add: numerals split add: nat_diff_split)
-
-
-subsubsection{*For @{term nat_case} and @{term nat_rec}*}
-
-lemma nat_case_number_of [simp]:
-     "nat_case a f (number_of v) =
-        (let pv = number_of (Int.pred v) in
-         if neg pv then a else f (nat pv))"
-by (simp split add: nat.split add: Let_def neg_number_of_pred_iff_0)
-
-lemma nat_case_add_eq_if [simp]:
-     "nat_case a f ((number_of v) + n) =
-       (let pv = number_of (Int.pred v) in
-         if neg pv then nat_case a f n else f (nat pv + n))"
-apply (subst add_eq_if)
-apply (simp split add: nat.split
-            del: nat_numeral_1_eq_1
-            add: nat_numeral_1_eq_1 [symmetric]
-                 numeral_1_eq_Suc_0 [symmetric]
-                 neg_number_of_pred_iff_0)
-done
-
-lemma nat_rec_number_of [simp]:
-     "nat_rec a f (number_of v) =
-        (let pv = number_of (Int.pred v) in
-         if neg pv then a else f (nat pv) (nat_rec a f (nat pv)))"
-apply (case_tac " (number_of v) ::nat")
-apply (simp_all (no_asm_simp) add: Let_def neg_number_of_pred_iff_0)
-apply (simp split add: split_if_asm)
-done
-
-lemma nat_rec_add_eq_if [simp]:
-     "nat_rec a f (number_of v + n) =
-        (let pv = number_of (Int.pred v) in
-         if neg pv then nat_rec a f n
-                   else f (nat pv + n) (nat_rec a f (nat pv + n)))"
-apply (subst add_eq_if)
-apply (simp split add: nat.split
-            del: nat_numeral_1_eq_1
-            add: nat_numeral_1_eq_1 [symmetric]
-                 numeral_1_eq_Suc_0 [symmetric]
-                 neg_number_of_pred_iff_0)
-done
-
-
-subsubsection{*Various Other Lemmas*}
-
-text {*Evens and Odds, for Mutilated Chess Board*}
-
-text{*Lemmas for specialist use, NOT as default simprules*}
-lemma nat_mult_2: "2 * z = (z+z::nat)"
-proof -
-  have "2*z = (1 + 1)*z" by simp
-  also have "... = z+z" by (simp add: left_distrib)
-  finally show ?thesis .
-qed
-
-lemma nat_mult_2_right: "z * 2 = (z+z::nat)"
-by (subst mult_commute, rule nat_mult_2)
-
-text{*Case analysis on @{term "n<2"}*}
-lemma less_2_cases: "(n::nat) < 2 ==> n = 0 | n = Suc 0"
-by arith
-
-lemma div2_Suc_Suc [simp]: "Suc(Suc m) div 2 = Suc (m div 2)"
-by arith
-
-lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)"
-by (simp add: nat_mult_2 [symmetric])
-
-lemma mod2_Suc_Suc [simp]: "Suc(Suc(m)) mod 2 = m mod 2"
-apply (subgoal_tac "m mod 2 < 2")
-apply (erule less_2_cases [THEN disjE])
-apply (simp_all (no_asm_simp) add: Let_def mod_Suc nat_1)
-done
-
-lemma mod2_gr_0 [simp]: "!!m::nat. (0 < m mod 2) = (m mod 2 = 1)"
-apply (subgoal_tac "m mod 2 < 2")
-apply (force simp del: mod_less_divisor, simp)
-done
-
-text{*Removal of Small Numerals: 0, 1 and (in additive positions) 2*}
-
-lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)"
-by simp
-
-lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)"
-by simp
-
-text{*Can be used to eliminate long strings of Sucs, but not by default*}
-lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n"
-by simp
-
-
-text{*These lemmas collapse some needless occurrences of Suc:
-    at least three Sucs, since two and fewer are rewritten back to Suc again!
-    We already have some rules to simplify operands smaller than 3.*}
-
-lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)"
-by (simp add: Suc3_eq_add_3)
-
-lemma mod_Suc_eq_mod_add3 [simp]: "m mod (Suc (Suc (Suc n))) = m mod (3+n)"
-by (simp add: Suc3_eq_add_3)
-
-lemma Suc_div_eq_add3_div: "(Suc (Suc (Suc m))) div n = (3+m) div n"
-by (simp add: Suc3_eq_add_3)
-
-lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n"
-by (simp add: Suc3_eq_add_3)
-
-lemmas Suc_div_eq_add3_div_number_of =
-    Suc_div_eq_add3_div [of _ "number_of v", standard]
-declare Suc_div_eq_add3_div_number_of [simp]
-
-lemmas Suc_mod_eq_add3_mod_number_of =
-    Suc_mod_eq_add3_mod [of _ "number_of v", standard]
-declare Suc_mod_eq_add3_mod_number_of [simp]
-
-
-subsubsection{*Special Simplification for Constants*}
-
-text{*These belong here, late in the development of HOL, to prevent their
-interfering with proofs of abstract properties of instances of the function
-@{term number_of}*}
-
-text{*These distributive laws move literals inside sums and differences.*}
-lemmas left_distrib_number_of = left_distrib [of _ _ "number_of v", standard]
-declare left_distrib_number_of [simp]
-
-lemmas right_distrib_number_of = right_distrib [of "number_of v", standard]
-declare right_distrib_number_of [simp]
-
-
-lemmas left_diff_distrib_number_of =
-    left_diff_distrib [of _ _ "number_of v", standard]
-declare left_diff_distrib_number_of [simp]
-
-lemmas right_diff_distrib_number_of =
-    right_diff_distrib [of "number_of v", standard]
-declare right_diff_distrib_number_of [simp]
-
-
-text{*These are actually for fields, like real: but where else to put them?*}
-lemmas zero_less_divide_iff_number_of =
-    zero_less_divide_iff [of "number_of w", standard]
-declare zero_less_divide_iff_number_of [simp,noatp]
-
-lemmas divide_less_0_iff_number_of =
-    divide_less_0_iff [of "number_of w", standard]
-declare divide_less_0_iff_number_of [simp,noatp]
-
-lemmas zero_le_divide_iff_number_of =
-    zero_le_divide_iff [of "number_of w", standard]
-declare zero_le_divide_iff_number_of [simp,noatp]
-
-lemmas divide_le_0_iff_number_of =
-    divide_le_0_iff [of "number_of w", standard]
-declare divide_le_0_iff_number_of [simp,noatp]
-
-
-(****
-IF times_divide_eq_right and times_divide_eq_left are removed as simprules,
-then these special-case declarations may be useful.
-
-text{*These simprules move numerals into numerators and denominators.*}
-lemma times_recip_eq_right [simp]: "a * (1/c) = a / (c::'a::field)"
-by (simp add: times_divide_eq)
-
-lemma times_recip_eq_left [simp]: "(1/c) * a = a / (c::'a::field)"
-by (simp add: times_divide_eq)
-
-lemmas times_divide_eq_right_number_of =
-    times_divide_eq_right [of "number_of w", standard]
-declare times_divide_eq_right_number_of [simp]
-
-lemmas times_divide_eq_right_number_of =
-    times_divide_eq_right [of _ _ "number_of w", standard]
-declare times_divide_eq_right_number_of [simp]
-
-lemmas times_divide_eq_left_number_of =
-    times_divide_eq_left [of _ "number_of w", standard]
-declare times_divide_eq_left_number_of [simp]
-
-lemmas times_divide_eq_left_number_of =
-    times_divide_eq_left [of _ _ "number_of w", standard]
-declare times_divide_eq_left_number_of [simp]
-
-****)
-
-text {*Replaces @{text "inverse #nn"} by @{text "1/#nn"}.  It looks
-  strange, but then other simprocs simplify the quotient.*}
-
-lemmas inverse_eq_divide_number_of =
-    inverse_eq_divide [of "number_of w", standard]
-declare inverse_eq_divide_number_of [simp]
-
-
-text {*These laws simplify inequalities, moving unary minus from a term
-into the literal.*}
-lemmas less_minus_iff_number_of =
-    less_minus_iff [of "number_of v", standard]
-declare less_minus_iff_number_of [simp,noatp]
-
-lemmas le_minus_iff_number_of =
-    le_minus_iff [of "number_of v", standard]
-declare le_minus_iff_number_of [simp,noatp]
-
-lemmas equation_minus_iff_number_of =
-    equation_minus_iff [of "number_of v", standard]
-declare equation_minus_iff_number_of [simp,noatp]
-
-
-lemmas minus_less_iff_number_of =
-    minus_less_iff [of _ "number_of v", standard]
-declare minus_less_iff_number_of [simp,noatp]
-
-lemmas minus_le_iff_number_of =
-    minus_le_iff [of _ "number_of v", standard]
-declare minus_le_iff_number_of [simp,noatp]
-
-lemmas minus_equation_iff_number_of =
-    minus_equation_iff [of _ "number_of v", standard]
-declare minus_equation_iff_number_of [simp,noatp]
-
-
-text{*To Simplify Inequalities Where One Side is the Constant 1*}
-
-lemma less_minus_iff_1 [simp,noatp]:
-  fixes b::"'b::{ordered_idom,number_ring}"
-  shows "(1 < - b) = (b < -1)"
-by auto
-
-lemma le_minus_iff_1 [simp,noatp]:
-  fixes b::"'b::{ordered_idom,number_ring}"
-  shows "(1 \<le> - b) = (b \<le> -1)"
-by auto
-
-lemma equation_minus_iff_1 [simp,noatp]:
-  fixes b::"'b::number_ring"
-  shows "(1 = - b) = (b = -1)"
-by (subst equation_minus_iff, auto)
-
-lemma minus_less_iff_1 [simp,noatp]:
-  fixes a::"'b::{ordered_idom,number_ring}"
-  shows "(- a < 1) = (-1 < a)"
-by auto
-
-lemma minus_le_iff_1 [simp,noatp]:
-  fixes a::"'b::{ordered_idom,number_ring}"
-  shows "(- a \<le> 1) = (-1 \<le> a)"
-by auto
-
-lemma minus_equation_iff_1 [simp,noatp]:
-  fixes a::"'b::number_ring"
-  shows "(- a = 1) = (a = -1)"
-by (subst minus_equation_iff, auto)
-
-
-text {*Cancellation of constant factors in comparisons (@{text "<"} and @{text "\<le>"}) *}
-
-lemmas mult_less_cancel_left_number_of =
-    mult_less_cancel_left [of "number_of v", standard]
-declare mult_less_cancel_left_number_of [simp,noatp]
-
-lemmas mult_less_cancel_right_number_of =
-    mult_less_cancel_right [of _ "number_of v", standard]
-declare mult_less_cancel_right_number_of [simp,noatp]
-
-lemmas mult_le_cancel_left_number_of =
-    mult_le_cancel_left [of "number_of v", standard]
-declare mult_le_cancel_left_number_of [simp,noatp]
-
-lemmas mult_le_cancel_right_number_of =
-    mult_le_cancel_right [of _ "number_of v", standard]
-declare mult_le_cancel_right_number_of [simp,noatp]
-
-
-text {*Multiplying out constant divisors in comparisons (@{text "<"}, @{text "\<le>"} and @{text "="}) *}
-
-lemmas le_divide_eq_number_of1 [simp] = le_divide_eq [of _ _ "number_of w", standard]
-lemmas divide_le_eq_number_of1 [simp] = divide_le_eq [of _ "number_of w", standard]
-lemmas less_divide_eq_number_of1 [simp] = less_divide_eq [of _ _ "number_of w", standard]
-lemmas divide_less_eq_number_of1 [simp] = divide_less_eq [of _ "number_of w", standard]
-lemmas eq_divide_eq_number_of1 [simp] = eq_divide_eq [of _ _ "number_of w", standard]
-lemmas divide_eq_eq_number_of1 [simp] = divide_eq_eq [of _ "number_of w", standard]
-
-
-subsubsection{*Optional Simplification Rules Involving Constants*}
-
-text{*Simplify quotients that are compared with a literal constant.*}
-
-lemmas le_divide_eq_number_of = le_divide_eq [of "number_of w", standard]
-lemmas divide_le_eq_number_of = divide_le_eq [of _ _ "number_of w", standard]
-lemmas less_divide_eq_number_of = less_divide_eq [of "number_of w", standard]
-lemmas divide_less_eq_number_of = divide_less_eq [of _ _ "number_of w", standard]
-lemmas eq_divide_eq_number_of = eq_divide_eq [of "number_of w", standard]
-lemmas divide_eq_eq_number_of = divide_eq_eq [of _ _ "number_of w", standard]
-
-
-text{*Not good as automatic simprules because they cause case splits.*}
-lemmas divide_const_simps =
-  le_divide_eq_number_of divide_le_eq_number_of less_divide_eq_number_of
-  divide_less_eq_number_of eq_divide_eq_number_of divide_eq_eq_number_of
-  le_divide_eq_1 divide_le_eq_1 less_divide_eq_1 divide_less_eq_1
-
-text{*Division By @{text "-1"}*}
-
-lemma divide_minus1 [simp]:
-     "x/-1 = -(x::'a::{field,division_by_zero,number_ring})"
-by simp
-
-lemma minus1_divide [simp]:
-     "-1 / (x::'a::{field,division_by_zero,number_ring}) = - (1/x)"
-by (simp add: divide_inverse inverse_minus_eq)
-
-lemma half_gt_zero_iff:
-     "(0 < r/2) = (0 < (r::'a::{ordered_field,division_by_zero,number_ring}))"
-by auto
-
-lemmas half_gt_zero = half_gt_zero_iff [THEN iffD2, standard]
-declare half_gt_zero [simp]
-
-(* The following lemma should appear in Divides.thy, but there the proof
-   doesn't work. *)
-
-lemma nat_dvd_not_less:
-  "[| 0 < m; m < n |] ==> \<not> n dvd (m::nat)"
-  by (unfold dvd_def) auto
-
-ML {*
-val divide_minus1 = @{thm divide_minus1};
-val minus1_divide = @{thm minus1_divide};
-*}
-
-end
--- a/src/HOL/Auth/Message.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Auth/Message.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Auth/Message
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1996  University of Cambridge
 
@@ -848,9 +847,9 @@
 (*Prove base case (subgoal i) and simplify others.  A typical base case
   concerns  Crypt K X \<notin> Key`shrK`bad  and cannot be proved by rewriting
   alone.*)
-fun prove_simple_subgoals_tac i = 
-    CLASIMPSET' (fn (cs, ss) => force_tac (cs, ss addsimps [@{thm image_eq_UN}])) i THEN
-    ALLGOALS (SIMPSET' asm_simp_tac)
+fun prove_simple_subgoals_tac (cs, ss) i = 
+    force_tac (cs, ss addsimps [@{thm image_eq_UN}]) i THEN
+    ALLGOALS (asm_simp_tac ss)
 
 (*Analysis of Fake cases.  Also works for messages that forward unknown parts,
   but this application is no longer necessary if analz_insert_eq is used.
@@ -875,8 +874,7 @@
 		  (cs addIs [@{thm analz_insertI},
 				   impOfSubs @{thm analz_subset_parts}]) 4 1))
 
-(*The explicit claset and simpset arguments help it work with Isar*)
-fun gen_spy_analz_tac (cs,ss) i =
+fun spy_analz_tac (cs,ss) i =
   DETERM
    (SELECT_GOAL
      (EVERY 
@@ -888,8 +886,6 @@
        REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])),
        DEPTH_SOLVE (atomic_spy_analz_tac (cs,ss) 1)]) i)
 
-val spy_analz_tac = CLASIMPSET' gen_spy_analz_tac;
-
 end
 *}
 
@@ -941,7 +937,7 @@
 lemmas Fake_parts_sing_imp_Un = Fake_parts_sing [THEN [2] rev_subsetD]
 
 method_setup spy_analz = {*
-    Scan.succeed (SIMPLE_METHOD' o Message.gen_spy_analz_tac o local_clasimpset_of) *}
+    Scan.succeed (SIMPLE_METHOD' o Message.spy_analz_tac o local_clasimpset_of) *}
     "for proving the Fake case when analz is involved"
 
 method_setup atomic_spy_analz = {*
--- a/src/HOL/Complex.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Complex.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -348,13 +348,13 @@
 
 subsection {* Completeness of the Complexes *}
 
-interpretation Re!: bounded_linear "Re"
+interpretation Re: bounded_linear "Re"
 apply (unfold_locales, simp, simp)
 apply (rule_tac x=1 in exI)
 apply (simp add: complex_norm_def)
 done
 
-interpretation Im!: bounded_linear "Im"
+interpretation Im: bounded_linear "Im"
 apply (unfold_locales, simp, simp)
 apply (rule_tac x=1 in exI)
 apply (simp add: complex_norm_def)
@@ -516,7 +516,7 @@
 lemma complex_mod_mult_cnj: "cmod (z * cnj z) = (cmod z)\<twosuperior>"
 by (simp add: norm_mult power2_eq_square)
 
-interpretation cnj!: bounded_linear "cnj"
+interpretation cnj: bounded_linear "cnj"
 apply (unfold_locales)
 apply (rule complex_cnj_add)
 apply (rule complex_cnj_scaleR)
--- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -6,7 +6,7 @@
   and a quantifier elimination procedure in Ferrante and Rackoff style *}
 
 theory Dense_Linear_Order
-imports Plain Groebner_Basis Main
+imports Main
 uses
   "~~/src/HOL/Tools/Qelim/langford_data.ML"
   "~~/src/HOL/Tools/Qelim/ferrante_rackoff_data.ML"
@@ -299,7 +299,7 @@
 *} "Langford's algorithm for quantifier elimination in dense linear orders"
 
 
-section {* Contructive dense linear orders yield QE for linear arithmetic over ordered Fields -- see @{text "Arith_Tools.thy"} *}
+section {* Contructive dense linear orders yield QE for linear arithmetic over ordered Fields *}
 
 text {* Linear order without upper bounds *}
 
@@ -637,7 +637,7 @@
   using eq_diff_eq[where a= x and b=t and c=0] by simp
 
 
-interpretation class_ordered_field_dense_linear_order!: constr_dense_linear_order
+interpretation class_ordered_field_dense_linear_order: constr_dense_linear_order
  "op <=" "op <"
    "\<lambda> x y. 1/2 * ((x::'a::{ordered_field,recpower,number_ring}) + y)"
 proof (unfold_locales, dlo, dlo, auto)
--- a/src/HOL/Decision_Procs/Ferrack.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Decision_Procs/Ferrack.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1995,6 +1995,8 @@
   "ferrack_test u = linrqe (A (A (Imp (Lt (Sub (Bound 1) (Bound 0)))
     (E (Eq (Sub (Add (Bound 0) (Bound 2)) (Bound 1)))))))"
 
+code_reserved SML oo
+
 ML {* @{code ferrack_test} () *}
 
 oracle linr_oracle = {*
--- a/src/HOL/Decision_Procs/MIR.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Decision_Procs/MIR.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -3783,8 +3783,7 @@
     also from mult_left_mono[OF xp] np have "?N s \<le> real n * x + ?N s" by simp
     finally have "?N (Floor s) \<le> real n * x + ?N s" .
     moreover
-    {from mult_strict_left_mono[OF x1] np 
-      have "real n *x + ?N s < real n + ?N s" by simp
+    {from x1 np have "real n *x + ?N s < real n + ?N s" by simp
       also from real_of_int_floor_add_one_gt[where r="?N s"] 
       have "\<dots> < real n + ?N (Floor s) + 1" by simp
       finally have "real n *x + ?N s < ?N (Floor s) + real (n+1)" by simp}
@@ -3809,8 +3808,7 @@
     moreover from mult_left_mono_neg[OF xp] nn have "?N s \<ge> real n * x + ?N s" by simp
     ultimately have "?N (Floor s) + 1 > real n * x + ?N s" by arith 
     moreover
-    {from mult_strict_left_mono_neg[OF x1, where c="real n"] nn
-      have "real n *x + ?N s \<ge> real n + ?N s" by simp 
+    {from x1 nn have "real n *x + ?N s \<ge> real n + ?N s" by simp
       moreover from real_of_int_floor_le[where r="?N s"]  have "real n + ?N s \<ge> real n + ?N (Floor s)" by simp
       ultimately have "real n *x + ?N s \<ge> ?N (Floor s) + real n" 
 	by (simp only: algebra_simps)}
--- a/src/HOL/Divides.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Divides.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -852,7 +852,7 @@
 
 text {* @{term "op dvd"} is a partial order *}
 
-interpretation dvd!: order "op dvd" "\<lambda>n m \<Colon> nat. n dvd m \<and> \<not> m dvd n"
+interpretation dvd: order "op dvd" "\<lambda>n m \<Colon> nat. n dvd m \<and> \<not> m dvd n"
   proof qed (auto intro: dvd_refl dvd_trans dvd_anti_sym)
 
 lemma nat_dvd_diff[simp]: "[| k dvd m; k dvd n |] ==> k dvd (m-n :: nat)"
@@ -1148,4 +1148,9 @@
   with j show ?thesis by blast
 qed
 
+lemma nat_dvd_not_less:
+  fixes m n :: nat
+  shows "0 < m \<Longrightarrow> m < n \<Longrightarrow> \<not> n dvd m"
+by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc)
+
 end
--- a/src/HOL/Finite_Set.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Finite_Set.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -928,7 +928,7 @@
 
 subsection {* Generalized summation over a set *}
 
-interpretation comm_monoid_add!: comm_monoid_mult "0::'a::comm_monoid_add" "op +"
+interpretation comm_monoid_add: comm_monoid_mult "0::'a::comm_monoid_add" "op +"
   proof qed (auto intro: add_assoc add_commute)
 
 definition setsum :: "('a => 'b) => 'a set => 'b::comm_monoid_add"
--- a/src/HOL/GCD.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/GCD.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -6,7 +6,7 @@
 header {* The Greatest Common Divisor *}
 
 theory GCD
-imports Plain Presburger Main
+imports Main
 begin
 
 text {*
--- a/src/HOL/Groebner_Basis.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Groebner_Basis.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Semiring normalization and Groebner Bases *}
 
 theory Groebner_Basis
-imports Arith_Tools
+imports NatBin
 uses
   "Tools/Groebner_Basis/misc.ML"
   "Tools/Groebner_Basis/normalizer_data.ML"
@@ -163,7 +163,7 @@
 
 end
 
-interpretation class_semiring!: gb_semiring
+interpretation class_semiring: gb_semiring
     "op +" "op *" "op ^" "0::'a::{comm_semiring_1, recpower}" "1"
   proof qed (auto simp add: algebra_simps power_Suc)
 
@@ -242,7 +242,7 @@
 end
 
 
-interpretation class_ring!: gb_ring "op +" "op *" "op ^"
+interpretation class_ring: gb_ring "op +" "op *" "op ^"
     "0::'a::{comm_semiring_1,recpower,number_ring}" 1 "op -" "uminus"
   proof qed simp_all
 
@@ -343,7 +343,7 @@
   thus "b = 0" by blast
 qed
 
-interpretation class_ringb!: ringb
+interpretation class_ringb: ringb
   "op +" "op *" "op ^" "0::'a::{idom,recpower,number_ring}" "1" "op -" "uminus"
 proof(unfold_locales, simp add: algebra_simps power_Suc, auto)
   fix w x y z ::"'a::{idom,recpower,number_ring}"
@@ -359,7 +359,7 @@
 
 declaration {* normalizer_funs @{thm class_ringb.ringb_axioms'} *}
 
-interpretation natgb!: semiringb
+interpretation natgb: semiringb
   "op +" "op *" "op ^" "0::nat" "1"
 proof (unfold_locales, simp add: algebra_simps power_Suc)
   fix w x y z ::"nat"
@@ -463,7 +463,7 @@
 
 subsection{* Groebner Bases for fields *}
 
-interpretation class_fieldgb!:
+interpretation class_fieldgb:
   fieldgb "op +" "op *" "op ^" "0::'a::{field,recpower,number_ring}" "1" "op -" "uminus" "op /" "inverse" apply (unfold_locales) by (simp_all add: divide_inverse)
 
 lemma divide_Numeral1: "(x::'a::{field,number_ring}) / Numeral1 = x" by simp
--- a/src/HOL/HOL.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/HOL.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1018,12 +1018,10 @@
   val contr_tac = Classical.contr_tac
   val dup_intr = Classical.dup_intr
   val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac
-  val claset = Classical.claset
   val rep_cs = Classical.rep_cs
   val cla_modifiers = Classical.cla_modifiers
   val cla_meth' = Classical.cla_meth'
 );
-val Blast_tac = Blast.Blast_tac;
 val blast_tac = Blast.blast_tac;
 *}
 
--- a/src/HOL/HahnBanach/Subspace.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/HahnBanach/Subspace.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -59,7 +59,7 @@
   assumes "vectorspace V"
   shows "0 \<in> U"
 proof -
-  interpret V!: vectorspace V by fact
+  interpret V: vectorspace V by fact
   have "U \<noteq> {}" by (rule non_empty)
   then obtain x where x: "x \<in> U" by blast
   then have "x \<in> V" .. then have "0 = x - x" by simp
--- a/src/HOL/HoareParallel/OG_Examples.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/HoareParallel/OG_Examples.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -443,7 +443,7 @@
 --{* 32 subgoals left *}
 apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
 
-apply(tactic {* TRYALL (simple_arith_tac @{context}) *})
+apply(tactic {* TRYALL (linear_arith_tac @{context}) *})
 --{* 9 subgoals left *}
 apply (force simp add:less_Suc_eq)
 apply(drule sym)
--- a/src/HOL/IMPP/Hoare.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/IMPP/Hoare.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/IMPP/Hoare.thy
-    ID:         $Id$
     Author:     David von Oheimb
     Copyright   1999 TUM
 *)
@@ -219,7 +218,7 @@
 apply           (rule hoare_derivs.conseq, intro strip, tactic "smp_tac 2 1", clarify, tactic "smp_tac 1 1",rule exI, rule exI, erule (1) conjI)
 prefer 7
 apply          (rule_tac hoare_derivs.Body, drule_tac spec, erule_tac mp, fast)
-apply         (tactic {* ALLGOALS (resolve_tac ((funpow 5 tl) (thms "hoare_derivs.intros")) THEN_ALL_NEW CLASET' fast_tac) *})
+apply         (tactic {* ALLGOALS (resolve_tac ((funpow 5 tl) (thms "hoare_derivs.intros")) THEN_ALL_NEW (fast_tac @{claset})) *})
 done
 
 lemma weak_Body: "G|-{P}. the (body pn) .{Q} ==> G|-{P}. BODY pn .{Q}"
@@ -291,7 +290,7 @@
    simp_tac @{simpset}, clarify_tac @{claset}, REPEAT o smp_tac 1]) *})
 apply       (simp_all (no_asm_use) add: triple_valid_def2)
 apply       (intro strip, tactic "smp_tac 2 1", blast) (* conseq *)
-apply      (tactic {* ALLGOALS (CLASIMPSET' clarsimp_tac) *}) (* Skip, Ass, Local *)
+apply      (tactic {* ALLGOALS (clarsimp_tac @{clasimpset}) *}) (* Skip, Ass, Local *)
 prefer 3 apply   (force) (* Call *)
 apply  (erule_tac [2] evaln_elim_cases) (* If *)
 apply   blast+
@@ -336,17 +335,17 @@
 lemma MGF_lemma1 [rule_format (no_asm)]: "state_not_singleton ==>  
   !pn:dom body. G|-{=}.BODY pn.{->} ==> WT c --> G|-{=}.c.{->}"
 apply (induct_tac c)
-apply        (tactic {* ALLGOALS (CLASIMPSET' clarsimp_tac) *})
+apply        (tactic {* ALLGOALS (clarsimp_tac @{clasimpset}) *})
 prefer 7 apply        (fast intro: domI)
 apply (erule_tac [6] MGT_alternD)
 apply       (unfold MGT_def)
 apply       (drule_tac [7] bspec, erule_tac [7] domI)
-apply       (rule_tac [7] escape, tactic {* CLASIMPSET' clarsimp_tac 7 *},
+apply       (rule_tac [7] escape, tactic {* clarsimp_tac @{clasimpset} 7 *},
   rule_tac [7] P1 = "%Z' s. s= (setlocs Z newlocs) [Loc Arg ::= fun Z]" in hoare_derivs.Call [THEN conseq1], erule_tac [7] conseq12)
 apply        (erule_tac [!] thin_rl)
 apply (rule hoare_derivs.Skip [THEN conseq2])
 apply (rule_tac [2] hoare_derivs.Ass [THEN conseq1])
-apply        (rule_tac [3] escape, tactic {* CLASIMPSET' clarsimp_tac 3 *},
+apply        (rule_tac [3] escape, tactic {* clarsimp_tac @{clasimpset} 3 *},
   rule_tac [3] P1 = "%Z' s. s= (Z[Loc loc::=fun Z])" in hoare_derivs.Local [THEN conseq1],
   erule_tac [3] conseq12)
 apply         (erule_tac [5] hoare_derivs.Comp, erule_tac [5] conseq12)
@@ -365,7 +364,7 @@
   shows "finite U ==> uG = mgt_call`U ==>  
   !G. G <= uG --> n <= card uG --> card G = card uG - n --> (!c. wt c --> P G {mgt c})"
 apply (induct_tac n)
-apply  (tactic {* ALLGOALS (CLASIMPSET' clarsimp_tac) *})
+apply  (tactic {* ALLGOALS (clarsimp_tac @{clasimpset}) *})
 apply  (subgoal_tac "G = mgt_call ` U")
 prefer 2
 apply   (simp add: card_seteq finite_imageI)
--- a/src/HOL/Imperative_HOL/Heap.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Imperative_HOL/Heap.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/Heap.thy
-    ID:         $Id$
     Author:     John Matthews, Galois Connections; Alexander Krauss, TU Muenchen
 *)
 
 header {* A polymorphic heap based on cantor encodings *}
 
 theory Heap
-imports Plain "~~/src/HOL/List" Countable Typerep
+imports Main Countable
 begin
 
 subsection {* Representable types *}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Imperative_HOL/Imperative_HOL_ex.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -0,0 +1,11 @@
+(*  Title:      HOL/Imperative_HOL/Imperative_HOL_ex.thy
+    Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
+*)
+
+header {* Mmonadic imperative HOL with examples *}
+
+theory Imperative_HOL_ex
+imports Imperative_HOL "ex/Imperative_Quicksort"
+begin
+
+end
--- a/src/HOL/Imperative_HOL/ROOT.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Imperative_HOL/ROOT.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -1,2 +1,2 @@
 
-use_thy "Imperative_HOL";
+use_thy "Imperative_HOL_ex";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -0,0 +1,639 @@
+(* Author: Lukas Bulwahn, TU Muenchen *)
+
+theory Imperative_Quicksort
+imports "~~/src/HOL/Imperative_HOL/Imperative_HOL" Subarray Multiset Efficient_Nat
+begin
+
+text {* We prove QuickSort correct in the Relational Calculus. *}
+
+definition swap :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap"
+where
+  "swap arr i j = (
+     do
+       x \<leftarrow> nth arr i;
+       y \<leftarrow> nth arr j;
+       upd i y arr;
+       upd j x arr;
+       return ()
+     done)"
+
+lemma swap_permutes:
+  assumes "crel (swap a i j) h h' rs"
+  shows "multiset_of (get_array a h') 
+  = multiset_of (get_array a h)"
+  using assms
+  unfolding swap_def
+  by (auto simp add: Heap.length_def multiset_of_swap dest: sym [of _ "h'"] elim!: crelE crel_nth crel_return crel_upd)
+
+function part1 :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat Heap"
+where
+  "part1 a left right p = (
+     if (right \<le> left) then return right
+     else (do
+       v \<leftarrow> nth a left;
+       (if (v \<le> p) then (part1 a (left + 1) right p)
+                    else (do swap a left right;
+  part1 a left (right - 1) p done))
+     done))"
+by pat_completeness auto
+
+termination
+by (relation "measure (\<lambda>(_,l,r,_). r - l )") auto
+
+declare part1.simps[simp del]
+
+lemma part_permutes:
+  assumes "crel (part1 a l r p) h h' rs"
+  shows "multiset_of (get_array a h') 
+  = multiset_of (get_array a h)"
+  using assms
+proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
+  case (1 a l r p h h' rs)
+  thus ?case
+    unfolding part1.simps [of a l r p]
+    by (elim crelE crel_if crel_return crel_nth) (auto simp add: swap_permutes)
+qed
+
+lemma part_returns_index_in_bounds:
+  assumes "crel (part1 a l r p) h h' rs"
+  assumes "l \<le> r"
+  shows "l \<le> rs \<and> rs \<le> r"
+using assms
+proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
+  case (1 a l r p h h' rs)
+  note cr = `crel (part1 a l r p) h h' rs`
+  show ?case
+  proof (cases "r \<le> l")
+    case True (* Terminating case *)
+    with cr `l \<le> r` show ?thesis
+      unfolding part1.simps[of a l r p]
+      by (elim crelE crel_if crel_return crel_nth) auto
+  next
+    case False (* recursive case *)
+    note rec_condition = this
+    let ?v = "get_array a h ! l"
+    show ?thesis
+    proof (cases "?v \<le> p")
+      case True
+      with cr False
+      have rec1: "crel (part1 a (l + 1) r p) h h' rs"
+        unfolding part1.simps[of a l r p]
+        by (elim crelE crel_nth crel_if crel_return) auto
+      from rec_condition have "l + 1 \<le> r" by arith
+      from 1(1)[OF rec_condition True rec1 `l + 1 \<le> r`]
+      show ?thesis by simp
+    next
+      case False
+      with rec_condition cr
+      obtain h1 where swp: "crel (swap a l r) h h1 ()"
+        and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
+        unfolding part1.simps[of a l r p]
+        by (elim crelE crel_nth crel_if crel_return) auto
+      from rec_condition have "l \<le> r - 1" by arith
+      from 1(2) [OF rec_condition False rec2 `l \<le> r - 1`] show ?thesis by fastsimp
+    qed
+  qed
+qed
+
+lemma part_length_remains:
+  assumes "crel (part1 a l r p) h h' rs"
+  shows "Heap.length a h = Heap.length a h'"
+using assms
+proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
+  case (1 a l r p h h' rs)
+  note cr = `crel (part1 a l r p) h h' rs`
+  
+  show ?case
+  proof (cases "r \<le> l")
+    case True (* Terminating case *)
+    with cr show ?thesis
+      unfolding part1.simps[of a l r p]
+      by (elim crelE crel_if crel_return crel_nth) auto
+  next
+    case False (* recursive case *)
+    with cr 1 show ?thesis
+      unfolding part1.simps [of a l r p] swap_def
+      by (auto elim!: crelE crel_if crel_nth crel_return crel_upd) fastsimp
+  qed
+qed
+
+lemma part_outer_remains:
+  assumes "crel (part1 a l r p) h h' rs"
+  shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
+  using assms
+proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
+  case (1 a l r p h h' rs)
+  note cr = `crel (part1 a l r p) h h' rs`
+  
+  show ?case
+  proof (cases "r \<le> l")
+    case True (* Terminating case *)
+    with cr show ?thesis
+      unfolding part1.simps[of a l r p]
+      by (elim crelE crel_if crel_return crel_nth) auto
+  next
+    case False (* recursive case *)
+    note rec_condition = this
+    let ?v = "get_array a h ! l"
+    show ?thesis
+    proof (cases "?v \<le> p")
+      case True
+      with cr False
+      have rec1: "crel (part1 a (l + 1) r p) h h' rs"
+        unfolding part1.simps[of a l r p]
+        by (elim crelE crel_nth crel_if crel_return) auto
+      from 1(1)[OF rec_condition True rec1]
+      show ?thesis by fastsimp
+    next
+      case False
+      with rec_condition cr
+      obtain h1 where swp: "crel (swap a l r) h h1 ()"
+        and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
+        unfolding part1.simps[of a l r p]
+        by (elim crelE crel_nth crel_if crel_return) auto
+      from swp rec_condition have
+        "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array a h ! i = get_array a h1 ! i"
+	unfolding swap_def
+	by (elim crelE crel_nth crel_upd crel_return) auto
+      with 1(2) [OF rec_condition False rec2] show ?thesis by fastsimp
+    qed
+  qed
+qed
+
+
+lemma part_partitions:
+  assumes "crel (part1 a l r p) h h' rs"
+  shows "(\<forall>i. l \<le> i \<and> i < rs \<longrightarrow> get_array (a::nat array) h' ! i \<le> p)
+  \<and> (\<forall>i. rs < i \<and> i \<le> r \<longrightarrow> get_array a h' ! i \<ge> p)"
+  using assms
+proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
+  case (1 a l r p h h' rs)
+  note cr = `crel (part1 a l r p) h h' rs`
+  
+  show ?case
+  proof (cases "r \<le> l")
+    case True (* Terminating case *)
+    with cr have "rs = r"
+      unfolding part1.simps[of a l r p]
+      by (elim crelE crel_if crel_return crel_nth) auto
+    with True
+    show ?thesis by auto
+  next
+    case False (* recursive case *)
+    note lr = this
+    let ?v = "get_array a h ! l"
+    show ?thesis
+    proof (cases "?v \<le> p")
+      case True
+      with lr cr
+      have rec1: "crel (part1 a (l + 1) r p) h h' rs"
+        unfolding part1.simps[of a l r p]
+        by (elim crelE crel_nth crel_if crel_return) auto
+      from True part_outer_remains[OF rec1] have a_l: "get_array a h' ! l \<le> p"
+	by fastsimp
+      have "\<forall>i. (l \<le> i = (l = i \<or> Suc l \<le> i))" by arith
+      with 1(1)[OF False True rec1] a_l show ?thesis
+	by auto
+    next
+      case False
+      with lr cr
+      obtain h1 where swp: "crel (swap a l r) h h1 ()"
+        and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
+        unfolding part1.simps[of a l r p]
+        by (elim crelE crel_nth crel_if crel_return) auto
+      from swp False have "get_array a h1 ! r \<ge> p"
+	unfolding swap_def
+	by (auto simp add: Heap.length_def elim!: crelE crel_nth crel_upd crel_return)
+      with part_outer_remains [OF rec2] lr have a_r: "get_array a h' ! r \<ge> p"
+	by fastsimp
+      have "\<forall>i. (i \<le> r = (i = r \<or> i \<le> r - 1))" by arith
+      with 1(2)[OF lr False rec2] a_r show ?thesis
+	by auto
+    qed
+  qed
+qed
+
+
+fun partition :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat Heap"
+where
+  "partition a left right = (do
+     pivot \<leftarrow> nth a right;
+     middle \<leftarrow> part1 a left (right - 1) pivot;
+     v \<leftarrow> nth a middle;
+     m \<leftarrow> return (if (v \<le> pivot) then (middle + 1) else middle);
+     swap a m right;
+     return m
+   done)"
+
+declare partition.simps[simp del]
+
+lemma partition_permutes:
+  assumes "crel (partition a l r) h h' rs"
+  shows "multiset_of (get_array a h') 
+  = multiset_of (get_array a h)"
+proof -
+    from assms part_permutes swap_permutes show ?thesis
+      unfolding partition.simps
+      by (elim crelE crel_return crel_nth crel_if crel_upd) auto
+qed
+
+lemma partition_length_remains:
+  assumes "crel (partition a l r) h h' rs"
+  shows "Heap.length a h = Heap.length a h'"
+proof -
+  from assms part_length_remains show ?thesis
+    unfolding partition.simps swap_def
+    by (elim crelE crel_return crel_nth crel_if crel_upd) auto
+qed
+
+lemma partition_outer_remains:
+  assumes "crel (partition a l r) h h' rs"
+  assumes "l < r"
+  shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
+proof -
+  from assms part_outer_remains part_returns_index_in_bounds show ?thesis
+    unfolding partition.simps swap_def
+    by (elim crelE crel_return crel_nth crel_if crel_upd) fastsimp
+qed
+
+lemma partition_returns_index_in_bounds:
+  assumes crel: "crel (partition a l r) h h' rs"
+  assumes "l < r"
+  shows "l \<le> rs \<and> rs \<le> r"
+proof -
+  from crel obtain middle h'' p where part: "crel (part1 a l (r - 1) p) h h'' middle"
+    and rs_equals: "rs = (if get_array a h'' ! middle \<le> get_array a h ! r then middle + 1
+         else middle)"
+    unfolding partition.simps
+    by (elim crelE crel_return crel_nth crel_if crel_upd) simp
+  from `l < r` have "l \<le> r - 1" by arith
+  from part_returns_index_in_bounds[OF part this] rs_equals `l < r` show ?thesis by auto
+qed
+
+lemma partition_partitions:
+  assumes crel: "crel (partition a l r) h h' rs"
+  assumes "l < r"
+  shows "(\<forall>i. l \<le> i \<and> i < rs \<longrightarrow> get_array (a::nat array) h' ! i \<le> get_array a h' ! rs) \<and>
+  (\<forall>i. rs < i \<and> i \<le> r \<longrightarrow> get_array a h' ! rs \<le> get_array a h' ! i)"
+proof -
+  let ?pivot = "get_array a h ! r" 
+  from crel obtain middle h1 where part: "crel (part1 a l (r - 1) ?pivot) h h1 middle"
+    and swap: "crel (swap a rs r) h1 h' ()"
+    and rs_equals: "rs = (if get_array a h1 ! middle \<le> ?pivot then middle + 1
+         else middle)"
+    unfolding partition.simps
+    by (elim crelE crel_return crel_nth crel_if crel_upd) simp
+  from swap have h'_def: "h' = Heap.upd a r (get_array a h1 ! rs)
+    (Heap.upd a rs (get_array a h1 ! r) h1)"
+    unfolding swap_def
+    by (elim crelE crel_return crel_nth crel_upd) simp
+  from swap have in_bounds: "r < Heap.length a h1 \<and> rs < Heap.length a h1"
+    unfolding swap_def
+    by (elim crelE crel_return crel_nth crel_upd) simp
+  from swap have swap_length_remains: "Heap.length a h1 = Heap.length a h'"
+    unfolding swap_def by (elim crelE crel_return crel_nth crel_upd) auto
+  from `l < r` have "l \<le> r - 1" by simp 
+  note middle_in_bounds = part_returns_index_in_bounds[OF part this]
+  from part_outer_remains[OF part] `l < r`
+  have "get_array a h ! r = get_array a h1 ! r"
+    by fastsimp
+  with swap
+  have right_remains: "get_array a h ! r = get_array a h' ! rs"
+    unfolding swap_def
+    by (auto simp add: Heap.length_def elim!: crelE crel_return crel_nth crel_upd) (cases "r = rs", auto)
+  from part_partitions [OF part]
+  show ?thesis
+  proof (cases "get_array a h1 ! middle \<le> ?pivot")
+    case True
+    with rs_equals have rs_equals: "rs = middle + 1" by simp
+    { 
+      fix i
+      assume i_is_left: "l \<le> i \<and> i < rs"
+      with swap_length_remains in_bounds middle_in_bounds rs_equals `l < r`
+      have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
+      from i_is_left rs_equals have "l \<le> i \<and> i < middle \<or> i = middle" by arith
+      with part_partitions[OF part] right_remains True
+      have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
+      with i_props h'_def in_bounds have "get_array a h' ! i \<le> get_array a h' ! rs"
+	unfolding Heap.upd_def Heap.length_def by simp
+    }
+    moreover
+    {
+      fix i
+      assume "rs < i \<and> i \<le> r"
+
+      hence "(rs < i \<and> i \<le> r - 1) \<or> (rs < i \<and> i = r)" by arith
+      hence "get_array a h' ! rs \<le> get_array a h' ! i"
+      proof
+	assume i_is: "rs < i \<and> i \<le> r - 1"
+	with swap_length_remains in_bounds middle_in_bounds rs_equals
+	have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
+	from part_partitions[OF part] rs_equals right_remains i_is
+	have "get_array a h' ! rs \<le> get_array a h1 ! i"
+	  by fastsimp
+	with i_props h'_def show ?thesis by fastsimp
+      next
+	assume i_is: "rs < i \<and> i = r"
+	with rs_equals have "Suc middle \<noteq> r" by arith
+	with middle_in_bounds `l < r` have "Suc middle \<le> r - 1" by arith
+	with part_partitions[OF part] right_remains 
+	have "get_array a h' ! rs \<le> get_array a h1 ! (Suc middle)"
+	  by fastsimp
+	with i_is True rs_equals right_remains h'_def
+	show ?thesis using in_bounds
+	  unfolding Heap.upd_def Heap.length_def
+	  by auto
+      qed
+    }
+    ultimately show ?thesis by auto
+  next
+    case False
+    with rs_equals have rs_equals: "middle = rs" by simp
+    { 
+      fix i
+      assume i_is_left: "l \<le> i \<and> i < rs"
+      with swap_length_remains in_bounds middle_in_bounds rs_equals
+      have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
+      from part_partitions[OF part] rs_equals right_remains i_is_left
+      have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
+      with i_props h'_def have "get_array a h' ! i \<le> get_array a h' ! rs"
+	unfolding Heap.upd_def by simp
+    }
+    moreover
+    {
+      fix i
+      assume "rs < i \<and> i \<le> r"
+      hence "(rs < i \<and> i \<le> r - 1) \<or> i = r" by arith
+      hence "get_array a h' ! rs \<le> get_array a h' ! i"
+      proof
+	assume i_is: "rs < i \<and> i \<le> r - 1"
+	with swap_length_remains in_bounds middle_in_bounds rs_equals
+	have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
+	from part_partitions[OF part] rs_equals right_remains i_is
+	have "get_array a h' ! rs \<le> get_array a h1 ! i"
+	  by fastsimp
+	with i_props h'_def show ?thesis by fastsimp
+      next
+	assume i_is: "i = r"
+	from i_is False rs_equals right_remains h'_def
+	show ?thesis using in_bounds
+	  unfolding Heap.upd_def Heap.length_def
+	  by auto
+      qed
+    }
+    ultimately
+    show ?thesis by auto
+  qed
+qed
+
+
+function quicksort :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap"
+where
+  "quicksort arr left right =
+     (if (right > left)  then
+        do
+          pivotNewIndex \<leftarrow> partition arr left right;
+          pivotNewIndex \<leftarrow> assert (\<lambda>x. left \<le> x \<and> x \<le> right) pivotNewIndex;
+          quicksort arr left (pivotNewIndex - 1);
+          quicksort arr (pivotNewIndex + 1) right
+        done
+     else return ())"
+by pat_completeness auto
+
+(* For termination, we must show that the pivotNewIndex is between left and right *) 
+termination
+by (relation "measure (\<lambda>(a, l, r). (r - l))") auto
+
+declare quicksort.simps[simp del]
+
+
+lemma quicksort_permutes:
+  assumes "crel (quicksort a l r) h h' rs"
+  shows "multiset_of (get_array a h') 
+  = multiset_of (get_array a h)"
+  using assms
+proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
+  case (1 a l r h h' rs)
+  with partition_permutes show ?case
+    unfolding quicksort.simps [of a l r]
+    by (elim crel_if crelE crel_assert crel_return) auto
+qed
+
+lemma length_remains:
+  assumes "crel (quicksort a l r) h h' rs"
+  shows "Heap.length a h = Heap.length a h'"
+using assms
+proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
+  case (1 a l r h h' rs)
+  with partition_length_remains show ?case
+    unfolding quicksort.simps [of a l r]
+    by (elim crel_if crelE crel_assert crel_return) auto
+qed
+
+lemma quicksort_outer_remains:
+  assumes "crel (quicksort a l r) h h' rs"
+   shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
+  using assms
+proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
+  case (1 a l r h h' rs)
+  note cr = `crel (quicksort a l r) h h' rs`
+  thus ?case
+  proof (cases "r > l")
+    case False
+    with cr have "h' = h"
+      unfolding quicksort.simps [of a l r]
+      by (elim crel_if crel_return) auto
+    thus ?thesis by simp
+  next
+  case True
+   { 
+      fix h1 h2 p ret1 ret2 i
+      assume part: "crel (partition a l r) h h1 p"
+      assume qs1: "crel (quicksort a l (p - 1)) h1 h2 ret1"
+      assume qs2: "crel (quicksort a (p + 1) r) h2 h' ret2"
+      assume pivot: "l \<le> p \<and> p \<le> r"
+      assume i_outer: "i < l \<or> r < i"
+      from  partition_outer_remains [OF part True] i_outer
+      have "get_array a h !i = get_array a h1 ! i" by fastsimp
+      moreover
+      with 1(1) [OF True pivot qs1] pivot i_outer
+      have "get_array a h1 ! i = get_array a h2 ! i" by auto
+      moreover
+      with qs2 1(2) [of p h2 h' ret2] True pivot i_outer
+      have "get_array a h2 ! i = get_array a h' ! i" by auto
+      ultimately have "get_array a h ! i= get_array a h' ! i" by simp
+    }
+    with cr show ?thesis
+      unfolding quicksort.simps [of a l r]
+      by (elim crel_if crelE crel_assert crel_return) auto
+  qed
+qed
+
+lemma quicksort_is_skip:
+  assumes "crel (quicksort a l r) h h' rs"
+  shows "r \<le> l \<longrightarrow> h = h'"
+  using assms
+  unfolding quicksort.simps [of a l r]
+  by (elim crel_if crel_return) auto
+ 
+lemma quicksort_sorts:
+  assumes "crel (quicksort a l r) h h' rs"
+  assumes l_r_length: "l < Heap.length a h" "r < Heap.length a h" 
+  shows "sorted (subarray l (r + 1) a h')"
+  using assms
+proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
+  case (1 a l r h h' rs)
+  note cr = `crel (quicksort a l r) h h' rs`
+  thus ?case
+  proof (cases "r > l")
+    case False
+    hence "l \<ge> r + 1 \<or> l = r" by arith 
+    with length_remains[OF cr] 1(5) show ?thesis
+      by (auto simp add: subarray_Nil subarray_single)
+  next
+    case True
+    { 
+      fix h1 h2 p
+      assume part: "crel (partition a l r) h h1 p"
+      assume qs1: "crel (quicksort a l (p - 1)) h1 h2 ()"
+      assume qs2: "crel (quicksort a (p + 1) r) h2 h' ()"
+      from partition_returns_index_in_bounds [OF part True]
+      have pivot: "l\<le> p \<and> p \<le> r" .
+     note length_remains = length_remains[OF qs2] length_remains[OF qs1] partition_length_remains[OF part]
+      from quicksort_outer_remains [OF qs2] quicksort_outer_remains [OF qs1] pivot quicksort_is_skip[OF qs1]
+      have pivot_unchanged: "get_array a h1 ! p = get_array a h' ! p" by (cases p, auto)
+        (*-- First of all, by induction hypothesis both sublists are sorted. *)
+      from 1(1)[OF True pivot qs1] length_remains pivot 1(5) 
+      have IH1: "sorted (subarray l p a h2)"  by (cases p, auto simp add: subarray_Nil)
+      from quicksort_outer_remains [OF qs2] length_remains
+      have left_subarray_remains: "subarray l p a h2 = subarray l p a h'"
+	by (simp add: subarray_eq_samelength_iff)
+      with IH1 have IH1': "sorted (subarray l p a h')" by simp
+      from 1(2)[OF True pivot qs2] pivot 1(5) length_remains
+      have IH2: "sorted (subarray (p + 1) (r + 1) a h')"
+        by (cases "Suc p \<le> r", auto simp add: subarray_Nil)
+           (* -- Secondly, both sublists remain partitioned. *)
+      from partition_partitions[OF part True]
+      have part_conds1: "\<forall>j. j \<in> set (subarray l p a h1) \<longrightarrow> j \<le> get_array a h1 ! p "
+        and part_conds2: "\<forall>j. j \<in> set (subarray (p + 1) (r + 1) a h1) \<longrightarrow> get_array a h1 ! p \<le> j"
+        by (auto simp add: all_in_set_subarray_conv)
+      from quicksort_outer_remains [OF qs1] quicksort_permutes [OF qs1] True
+        length_remains 1(5) pivot multiset_of_sublist [of l p "get_array a h1" "get_array a h2"]
+      have multiset_partconds1: "multiset_of (subarray l p a h2) = multiset_of (subarray l p a h1)"
+	unfolding Heap.length_def subarray_def by (cases p, auto)
+      with left_subarray_remains part_conds1 pivot_unchanged
+      have part_conds2': "\<forall>j. j \<in> set (subarray l p a h') \<longrightarrow> j \<le> get_array a h' ! p"
+        by (simp, subst set_of_multiset_of[symmetric], simp)
+          (* -- These steps are the analogous for the right sublist \<dots> *)
+      from quicksort_outer_remains [OF qs1] length_remains
+      have right_subarray_remains: "subarray (p + 1) (r + 1) a h1 = subarray (p + 1) (r + 1) a h2"
+	by (auto simp add: subarray_eq_samelength_iff)
+      from quicksort_outer_remains [OF qs2] quicksort_permutes [OF qs2] True
+        length_remains 1(5) pivot multiset_of_sublist [of "p + 1" "r + 1" "get_array a h2" "get_array a h'"]
+      have multiset_partconds2: "multiset_of (subarray (p + 1) (r + 1) a h') = multiset_of (subarray (p + 1) (r + 1) a h2)"
+        unfolding Heap.length_def subarray_def by auto
+      with right_subarray_remains part_conds2 pivot_unchanged
+      have part_conds1': "\<forall>j. j \<in> set (subarray (p + 1) (r + 1) a h') \<longrightarrow> get_array a h' ! p \<le> j"
+        by (simp, subst set_of_multiset_of[symmetric], simp)
+          (* -- Thirdly and finally, we show that the array is sorted
+          following from the facts above. *)
+      from True pivot 1(5) length_remains have "subarray l (r + 1) a h' = subarray l p a h' @ [get_array a h' ! p] @ subarray (p + 1) (r + 1) a h'"
+	by (simp add: subarray_nth_array_Cons, cases "l < p") (auto simp add: subarray_append subarray_Nil)
+      with IH1' IH2 part_conds1' part_conds2' pivot have ?thesis
+	unfolding subarray_def
+	apply (auto simp add: sorted_append sorted_Cons all_in_set_sublist'_conv)
+	by (auto simp add: set_sublist' dest: le_trans [of _ "get_array a h' ! p"])
+    }
+    with True cr show ?thesis
+      unfolding quicksort.simps [of a l r]
+      by (elim crel_if crel_return crelE crel_assert) auto
+  qed
+qed
+
+
+lemma quicksort_is_sort:
+  assumes crel: "crel (quicksort a 0 (Heap.length a h - 1)) h h' rs"
+  shows "get_array a h' = sort (get_array a h)"
+proof (cases "get_array a h = []")
+  case True
+  with quicksort_is_skip[OF crel] show ?thesis
+  unfolding Heap.length_def by simp
+next
+  case False
+  from quicksort_sorts [OF crel] False have "sorted (sublist' 0 (List.length (get_array a h)) (get_array a h'))"
+    unfolding Heap.length_def subarray_def by auto
+  with length_remains[OF crel] have "sorted (get_array a h')"
+    unfolding Heap.length_def by simp
+  with quicksort_permutes [OF crel] properties_for_sort show ?thesis by fastsimp
+qed
+
+subsection {* No Errors in quicksort *}
+text {* We have proved that quicksort sorts (if no exceptions occur).
+We will now show that exceptions do not occur. *}
+
+lemma noError_part1: 
+  assumes "l < Heap.length a h" "r < Heap.length a h"
+  shows "noError (part1 a l r p) h"
+  using assms
+proof (induct a l r p arbitrary: h rule: part1.induct)
+  case (1 a l r p)
+  thus ?case
+    unfolding part1.simps [of a l r] swap_def
+    by (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd elim!: crelE crel_upd crel_nth crel_return)
+qed
+
+lemma noError_partition:
+  assumes "l < r" "l < Heap.length a h" "r < Heap.length a h"
+  shows "noError (partition a l r) h"
+using assms
+unfolding partition.simps swap_def
+apply (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd noError_part1 elim!: crelE crel_upd crel_nth crel_return)
+apply (frule part_length_remains)
+apply (frule part_returns_index_in_bounds)
+apply auto
+apply (frule part_length_remains)
+apply (frule part_returns_index_in_bounds)
+apply auto
+apply (frule part_length_remains)
+apply auto
+done
+
+lemma noError_quicksort:
+  assumes "l < Heap.length a h" "r < Heap.length a h"
+  shows "noError (quicksort a l r) h"
+using assms
+proof (induct a l r arbitrary: h rule: quicksort.induct)
+  case (1 a l ri h)
+  thus ?case
+    unfolding quicksort.simps [of a l ri]
+    apply (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd noError_assert noError_partition)
+    apply (frule partition_returns_index_in_bounds)
+    apply auto
+    apply (frule partition_returns_index_in_bounds)
+    apply auto
+    apply (auto elim!: crel_assert dest!: partition_length_remains length_remains)
+    apply (subgoal_tac "Suc r \<le> ri \<or> r = ri") 
+    apply (erule disjE)
+    apply auto
+    unfolding quicksort.simps [of a "Suc ri" ri]
+    apply (auto intro!: noError_if noError_return)
+    done
+qed
+
+
+subsection {* Example *}
+
+definition "qsort a = do
+    k \<leftarrow> length a;
+    quicksort a 0 (k - 1);
+    return a
+  done"
+
+ML {* @{code qsort} (Array.fromList [42, 2, 3, 5, 0, 1705, 8, 3, 15]) () *}
+
+export_code qsort in SML_imp module_name QSort
+export_code qsort in OCaml module_name QSort file -
+export_code qsort in OCaml_imp module_name QSort file -
+export_code qsort in Haskell module_name QSort file -
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Imperative_HOL/ex/Subarray.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -0,0 +1,66 @@
+theory Subarray
+imports Array Sublist
+begin
+
+definition subarray :: "nat \<Rightarrow> nat \<Rightarrow> ('a::heap) array \<Rightarrow> heap \<Rightarrow> 'a list"
+where
+  "subarray n m a h \<equiv> sublist' n m (get_array a h)"
+
+lemma subarray_upd: "i \<ge> m \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
+apply (simp add: subarray_def Heap.upd_def)
+apply (simp add: sublist'_update1)
+done
+
+lemma subarray_upd2: " i < n  \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
+apply (simp add: subarray_def Heap.upd_def)
+apply (subst sublist'_update2)
+apply fastsimp
+apply simp
+done
+
+lemma subarray_upd3: "\<lbrakk> n \<le> i; i < m\<rbrakk> \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h[i - n := v]"
+unfolding subarray_def Heap.upd_def
+by (simp add: sublist'_update3)
+
+lemma subarray_Nil: "n \<ge> m \<Longrightarrow> subarray n m a h = []"
+by (simp add: subarray_def sublist'_Nil')
+
+lemma subarray_single: "\<lbrakk> n < Heap.length a h \<rbrakk> \<Longrightarrow> subarray n (Suc n) a h = [get_array a h ! n]" 
+by (simp add: subarray_def Heap.length_def sublist'_single)
+
+lemma length_subarray: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray n m a h) = m - n"
+by (simp add: subarray_def Heap.length_def length_sublist')
+
+lemma length_subarray_0: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray 0 m a h) = m"
+by (simp add: length_subarray)
+
+lemma subarray_nth_array_Cons: "\<lbrakk> i < Heap.length a h; i < j \<rbrakk> \<Longrightarrow> (get_array a h ! i) # subarray (Suc i) j a h = subarray i j a h"
+unfolding Heap.length_def subarray_def
+by (simp add: sublist'_front)
+
+lemma subarray_nth_array_back: "\<lbrakk> i < j; j \<le> Heap.length a h\<rbrakk> \<Longrightarrow> subarray i j a h = subarray i (j - 1) a h @ [get_array a h ! (j - 1)]"
+unfolding Heap.length_def subarray_def
+by (simp add: sublist'_back)
+
+lemma subarray_append: "\<lbrakk> i < j; j < k \<rbrakk> \<Longrightarrow> subarray i j a h @ subarray j k a h = subarray i k a h"
+unfolding subarray_def
+by (simp add: sublist'_append)
+
+lemma subarray_all: "subarray 0 (Heap.length a h) a h = get_array a h"
+unfolding Heap.length_def subarray_def
+by (simp add: sublist'_all)
+
+lemma nth_subarray: "\<lbrakk> k < j - i; j \<le> Heap.length a h \<rbrakk> \<Longrightarrow> subarray i j a h ! k = get_array a h ! (i + k)"
+unfolding Heap.length_def subarray_def
+by (simp add: nth_sublist')
+
+lemma subarray_eq_samelength_iff: "Heap.length a h = Heap.length a h' \<Longrightarrow> (subarray i j a h = subarray i j a h') = (\<forall>i'. i \<le> i' \<and> i' < j \<longrightarrow> get_array a h ! i' = get_array a h' ! i')"
+unfolding Heap.length_def subarray_def by (rule sublist'_eq_samelength_iff)
+
+lemma all_in_set_subarray_conv: "(\<forall>j. j \<in> set (subarray l r a h) \<longrightarrow> P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
+unfolding subarray_def Heap.length_def by (rule all_in_set_sublist'_conv)
+
+lemma ball_in_set_subarray_conv: "(\<forall>j \<in> set (subarray l r a h). P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
+unfolding subarray_def Heap.length_def by (rule ball_in_set_sublist'_conv)
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Imperative_HOL/ex/Sublist.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -0,0 +1,505 @@
+(* $Id$ *)
+
+header {* Slices of lists *}
+
+theory Sublist
+imports Multiset
+begin
+
+
+lemma sublist_split: "i \<le> j \<and> j \<le> k \<Longrightarrow> sublist xs {i..<j} @ sublist xs {j..<k} = sublist xs {i..<k}" 
+apply (induct xs arbitrary: i j k)
+apply simp
+apply (simp only: sublist_Cons)
+apply simp
+apply safe
+apply simp
+apply (erule_tac x="0" in meta_allE)
+apply (erule_tac x="j - 1" in meta_allE)
+apply (erule_tac x="k - 1" in meta_allE)
+apply (subgoal_tac "0 \<le> j - 1 \<and> j - 1 \<le> k - 1")
+apply simp
+apply (subgoal_tac "{ja. Suc ja < j} = {0..<j - Suc 0}")
+apply (subgoal_tac "{ja. j \<le> Suc ja \<and> Suc ja < k} = {j - Suc 0..<k - Suc 0}")
+apply (subgoal_tac "{j. Suc j < k} = {0..<k - Suc 0}")
+apply simp
+apply fastsimp
+apply fastsimp
+apply fastsimp
+apply fastsimp
+apply (erule_tac x="i - 1" in meta_allE)
+apply (erule_tac x="j - 1" in meta_allE)
+apply (erule_tac x="k - 1" in meta_allE)
+apply (subgoal_tac " {ja. i \<le> Suc ja \<and> Suc ja < j} = {i - 1 ..<j - 1}")
+apply (subgoal_tac " {ja. j \<le> Suc ja \<and> Suc ja < k} = {j - 1..<k - 1}")
+apply (subgoal_tac "{j. i \<le> Suc j \<and> Suc j < k} = {i - 1..<k - 1}")
+apply (subgoal_tac " i - 1 \<le> j - 1 \<and> j - 1 \<le> k - 1")
+apply simp
+apply fastsimp
+apply fastsimp
+apply fastsimp
+apply fastsimp
+done
+
+lemma sublist_update1: "i \<notin> inds \<Longrightarrow> sublist (xs[i := v]) inds = sublist xs inds"
+apply (induct xs arbitrary: i inds)
+apply simp
+apply (case_tac i)
+apply (simp add: sublist_Cons)
+apply (simp add: sublist_Cons)
+done
+
+lemma sublist_update2: "i \<in> inds \<Longrightarrow> sublist (xs[i := v]) inds = (sublist xs inds)[(card {k \<in> inds. k < i}):= v]"
+proof (induct xs arbitrary: i inds)
+  case Nil thus ?case by simp
+next
+  case (Cons x xs)
+  thus ?case
+  proof (cases i)
+    case 0 with Cons show ?thesis by (simp add: sublist_Cons)
+  next
+    case (Suc i')
+    with Cons show ?thesis
+      apply simp
+      apply (simp add: sublist_Cons)
+      apply auto
+      apply (auto simp add: nat.split)
+      apply (simp add: card_less_Suc[symmetric])
+      apply (simp add: card_less_Suc2)
+      done
+  qed
+qed
+
+lemma sublist_update: "sublist (xs[i := v]) inds = (if i \<in> inds then (sublist xs inds)[(card {k \<in> inds. k < i}) := v] else sublist xs inds)"
+by (simp add: sublist_update1 sublist_update2)
+
+lemma sublist_take: "sublist xs {j. j < m} = take m xs"
+apply (induct xs arbitrary: m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (simp add: sublist_Cons)
+done
+
+lemma sublist_take': "sublist xs {0..<m} = take m xs"
+apply (induct xs arbitrary: m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (simp add: sublist_Cons sublist_take)
+done
+
+lemma sublist_all[simp]: "sublist xs {j. j < length xs} = xs"
+apply (induct xs)
+apply simp
+apply (simp add: sublist_Cons)
+done
+
+lemma sublist_all'[simp]: "sublist xs {0..<length xs} = xs"
+apply (induct xs)
+apply simp
+apply (simp add: sublist_Cons)
+done
+
+lemma sublist_single: "a < length xs \<Longrightarrow> sublist xs {a} = [xs ! a]"
+apply (induct xs arbitrary: a)
+apply simp
+apply(case_tac aa)
+apply simp
+apply (simp add: sublist_Cons)
+apply simp
+apply (simp add: sublist_Cons)
+done
+
+lemma sublist_is_Nil: "\<forall>i \<in> inds. i \<ge> length xs \<Longrightarrow> sublist xs inds = []" 
+apply (induct xs arbitrary: inds)
+apply simp
+apply (simp add: sublist_Cons)
+apply auto
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply auto
+done
+
+lemma sublist_Nil': "sublist xs inds = [] \<Longrightarrow> \<forall>i \<in> inds. i \<ge> length xs"
+apply (induct xs arbitrary: inds)
+apply simp
+apply (simp add: sublist_Cons)
+apply (auto split: if_splits)
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply (case_tac x, auto)
+done
+
+lemma sublist_Nil[simp]: "(sublist xs inds = []) = (\<forall>i \<in> inds. i \<ge> length xs)"
+apply (induct xs arbitrary: inds)
+apply simp
+apply (simp add: sublist_Cons)
+apply auto
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply (case_tac x, auto)
+done
+
+lemma sublist_eq_subseteq: " \<lbrakk> inds' \<subseteq> inds; sublist xs inds = sublist ys inds \<rbrakk> \<Longrightarrow> sublist xs inds' = sublist ys inds'"
+apply (induct xs arbitrary: ys inds inds')
+apply simp
+apply (drule sym, rule sym)
+apply (simp add: sublist_Nil, fastsimp)
+apply (case_tac ys)
+apply (simp add: sublist_Nil, fastsimp)
+apply (auto simp add: sublist_Cons)
+apply (erule_tac x="list" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds'}" in meta_allE)
+apply fastsimp
+apply (erule_tac x="list" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds'}" in meta_allE)
+apply fastsimp
+done
+
+lemma sublist_eq: "\<lbrakk> \<forall>i \<in> inds. ((i < length xs) \<and> (i < length ys)) \<or> ((i \<ge> length xs ) \<and> (i \<ge> length ys)); \<forall>i \<in> inds. xs ! i = ys ! i \<rbrakk> \<Longrightarrow> sublist xs inds = sublist ys inds"
+apply (induct xs arbitrary: ys inds)
+apply simp
+apply (rule sym, simp add: sublist_Nil)
+apply (case_tac ys)
+apply (simp add: sublist_Nil)
+apply (auto simp add: sublist_Cons)
+apply (erule_tac x="list" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply fastsimp
+apply (erule_tac x="list" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply fastsimp
+done
+
+lemma sublist_eq_samelength: "\<lbrakk> length xs = length ys; \<forall>i \<in> inds. xs ! i = ys ! i \<rbrakk> \<Longrightarrow> sublist xs inds = sublist ys inds"
+by (rule sublist_eq, auto)
+
+lemma sublist_eq_samelength_iff: "length xs = length ys \<Longrightarrow> (sublist xs inds = sublist ys inds) = (\<forall>i \<in> inds. xs ! i = ys ! i)"
+apply (induct xs arbitrary: ys inds)
+apply simp
+apply (rule sym, simp add: sublist_Nil)
+apply (case_tac ys)
+apply (simp add: sublist_Nil)
+apply (auto simp add: sublist_Cons)
+apply (case_tac i)
+apply auto
+apply (case_tac i)
+apply auto
+done
+
+section {* Another sublist function *}
+
+function sublist' :: "nat \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
+where
+  "sublist' n m [] = []"
+| "sublist' n 0 xs = []"
+| "sublist' 0 (Suc m) (x#xs) = (x#sublist' 0 m xs)"
+| "sublist' (Suc n) (Suc m) (x#xs) = sublist' n m xs"
+by pat_completeness auto
+termination by lexicographic_order
+
+subsection {* Proving equivalence to the other sublist command *}
+
+lemma sublist'_sublist: "sublist' n m xs = sublist xs {j. n \<le> j \<and> j < m}"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac n)
+apply (case_tac m)
+apply simp
+apply (simp add: sublist_Cons)
+apply (case_tac m)
+apply simp
+apply (simp add: sublist_Cons)
+done
+
+
+lemma "sublist' n m xs = sublist xs {n..<m}"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac n, case_tac m)
+apply simp
+apply simp
+apply (simp add: sublist_take')
+apply (case_tac m)
+apply simp
+apply (simp add: sublist_Cons sublist'_sublist)
+done
+
+
+subsection {* Showing equivalence to use of drop and take for definition *}
+
+lemma "sublist' n m xs = take (m - n) (drop n xs)"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (case_tac n)
+apply simp
+apply simp
+done
+
+subsection {* General lemma about sublist *}
+
+lemma sublist'_Nil[simp]: "sublist' i j [] = []"
+by simp
+
+lemma sublist'_Cons[simp]: "sublist' i (Suc j) (x#xs) = (case i of 0 \<Rightarrow> (x # sublist' 0 j xs) | Suc i' \<Rightarrow>  sublist' i' j xs)"
+by (cases i) auto
+
+lemma sublist'_Cons2[simp]: "sublist' i j (x#xs) = (if (j = 0) then [] else ((if (i = 0) then [x] else []) @ sublist' (i - 1) (j - 1) xs))"
+apply (cases j)
+apply auto
+apply (cases i)
+apply auto
+done
+
+lemma sublist_n_0: "sublist' n 0 xs = []"
+by (induct xs, auto)
+
+lemma sublist'_Nil': "n \<ge> m \<Longrightarrow> sublist' n m xs = []"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (case_tac n)
+apply simp
+apply simp
+done
+
+lemma sublist'_Nil2: "n \<ge> length xs \<Longrightarrow> sublist' n m xs = []"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (case_tac n)
+apply simp
+apply simp
+done
+
+lemma sublist'_Nil3: "(sublist' n m xs = []) = ((n \<ge> m) \<or> (n \<ge> length xs))"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (case_tac n)
+apply simp
+apply simp
+done
+
+lemma sublist'_notNil: "\<lbrakk> n < length xs; n < m \<rbrakk> \<Longrightarrow> sublist' n m xs \<noteq> []"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (case_tac n)
+apply simp
+apply simp
+done
+
+lemma sublist'_single: "n < length xs \<Longrightarrow> sublist' n (Suc n) xs = [xs ! n]"
+apply (induct xs arbitrary: n)
+apply simp
+apply simp
+apply (case_tac n)
+apply (simp add: sublist_n_0)
+apply simp
+done
+
+lemma sublist'_update1: "i \<ge> m \<Longrightarrow> sublist' n m (xs[i:=v]) = sublist' n m xs"
+apply (induct xs arbitrary: n m i)
+apply simp
+apply simp
+apply (case_tac i)
+apply simp
+apply simp
+done
+
+lemma sublist'_update2: "i < n \<Longrightarrow> sublist' n m (xs[i:=v]) = sublist' n m xs"
+apply (induct xs arbitrary: n m i)
+apply simp
+apply simp
+apply (case_tac i)
+apply simp
+apply simp
+done
+
+lemma sublist'_update3: "\<lbrakk>n \<le> i; i < m\<rbrakk> \<Longrightarrow> sublist' n m (xs[i := v]) = (sublist' n m xs)[i - n := v]"
+proof (induct xs arbitrary: n m i)
+  case Nil thus ?case by auto
+next
+  case (Cons x xs)
+  thus ?case
+    apply -
+    apply auto
+    apply (cases i)
+    apply auto
+    apply (cases i)
+    apply auto
+    done
+qed
+
+lemma "\<lbrakk> sublist' i j xs = sublist' i j ys; n \<ge> i; m \<le> j \<rbrakk> \<Longrightarrow> sublist' n m xs = sublist' n m ys"
+proof (induct xs arbitrary: i j ys n m)
+  case Nil
+  thus ?case
+    apply -
+    apply (rule sym, drule sym)
+    apply (simp add: sublist'_Nil)
+    apply (simp add: sublist'_Nil3)
+    apply arith
+    done
+next
+  case (Cons x xs i j ys n m)
+  note c = this
+  thus ?case
+  proof (cases m)
+    case 0 thus ?thesis by (simp add: sublist_n_0)
+  next
+    case (Suc m')
+    note a = this
+    thus ?thesis
+    proof (cases n)
+      case 0 note b = this
+      show ?thesis
+      proof (cases ys)
+	case Nil  with a b Cons.prems show ?thesis by (simp add: sublist'_Nil3)
+      next
+	case (Cons y ys)
+	show ?thesis
+	proof (cases j)
+	  case 0 with a b Cons.prems show ?thesis by simp
+	next
+	  case (Suc j') with a b Cons.prems Cons show ?thesis 
+	    apply -
+	    apply (simp, rule Cons.hyps [of "0" "j'" "ys" "0" "m'"], auto)
+	    done
+	qed
+      qed
+    next
+      case (Suc n')
+      show ?thesis
+      proof (cases ys)
+	case Nil with Suc a Cons.prems show ?thesis by (auto simp add: sublist'_Nil3)
+      next
+	case (Cons y ys) with Suc a Cons.prems show ?thesis
+	  apply -
+	  apply simp
+	  apply (cases j)
+	  apply simp
+	  apply (cases i)
+	  apply simp
+	  apply (rule_tac j="nat" in Cons.hyps [of "0" _ "ys" "n'" "m'"])
+	  apply simp
+	  apply simp
+	  apply simp
+	  apply simp
+	  apply (rule_tac i="nata" and j="nat" in Cons.hyps [of _ _ "ys" "n'" "m'"])
+	  apply simp
+	  apply simp
+	  apply simp
+	  done
+      qed
+    qed
+  qed
+qed
+
+lemma length_sublist': "j \<le> length xs \<Longrightarrow> length (sublist' i j xs) = j - i"
+by (induct xs arbitrary: i j, auto)
+
+lemma sublist'_front: "\<lbrakk> i < j; i < length xs \<rbrakk> \<Longrightarrow> sublist' i j xs = xs ! i # sublist' (Suc i) j xs"
+apply (induct xs arbitrary: a i j)
+apply simp
+apply (case_tac j)
+apply simp
+apply (case_tac i)
+apply simp
+apply simp
+done
+
+lemma sublist'_back: "\<lbrakk> i < j; j \<le> length xs \<rbrakk> \<Longrightarrow> sublist' i j xs = sublist' i (j - 1) xs @ [xs ! (j - 1)]"
+apply (induct xs arbitrary: a i j)
+apply simp
+apply simp
+apply (case_tac j)
+apply simp
+apply auto
+apply (case_tac nat)
+apply auto
+done
+
+(* suffices that j \<le> length xs and length ys *) 
+lemma sublist'_eq_samelength_iff: "length xs = length ys \<Longrightarrow> (sublist' i j xs  = sublist' i j ys) = (\<forall>i'. i \<le> i' \<and> i' < j \<longrightarrow> xs ! i' = ys ! i')"
+proof (induct xs arbitrary: ys i j)
+  case Nil thus ?case by simp
+next
+  case (Cons x xs)
+  thus ?case
+    apply -
+    apply (cases ys)
+    apply simp
+    apply simp
+    apply auto
+    apply (case_tac i', auto)
+    apply (erule_tac x="Suc i'" in allE, auto)
+    apply (erule_tac x="i' - 1" in allE, auto)
+    apply (case_tac i', auto)
+    apply (erule_tac x="Suc i'" in allE, auto)
+    done
+qed
+
+lemma sublist'_all[simp]: "sublist' 0 (length xs) xs = xs"
+by (induct xs, auto)
+
+lemma sublist'_sublist': "sublist' n m (sublist' i j xs) = sublist' (i + n) (min (i + m) j) xs" 
+by (induct xs arbitrary: i j n m) (auto simp add: min_diff)
+
+lemma sublist'_append: "\<lbrakk> i \<le> j; j \<le> k \<rbrakk> \<Longrightarrow>(sublist' i j xs) @ (sublist' j k xs) = sublist' i k xs"
+by (induct xs arbitrary: i j k) auto
+
+lemma nth_sublist': "\<lbrakk> k < j - i; j \<le> length xs \<rbrakk> \<Longrightarrow> (sublist' i j xs) ! k = xs ! (i + k)"
+apply (induct xs arbitrary: i j k)
+apply auto
+apply (case_tac k)
+apply auto
+apply (case_tac i)
+apply auto
+done
+
+lemma set_sublist': "set (sublist' i j xs) = {x. \<exists>k. i \<le> k \<and> k < j \<and> k < List.length xs \<and> x = xs ! k}"
+apply (simp add: sublist'_sublist)
+apply (simp add: set_sublist)
+apply auto
+done
+
+lemma all_in_set_sublist'_conv: "(\<forall>j. j \<in> set (sublist' l r xs) \<longrightarrow> P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < List.length xs \<longrightarrow> P (xs ! k))"
+unfolding set_sublist' by blast
+
+lemma ball_in_set_sublist'_conv: "(\<forall>j \<in> set (sublist' l r xs). P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < List.length xs \<longrightarrow> P (xs ! k))"
+unfolding set_sublist' by blast
+
+
+lemma multiset_of_sublist:
+assumes l_r: "l \<le> r \<and> r \<le> List.length xs"
+assumes left: "\<forall> i. i < l \<longrightarrow> (xs::'a list) ! i = ys ! i"
+assumes right: "\<forall> i. i \<ge> r \<longrightarrow> (xs::'a list) ! i = ys ! i"
+assumes multiset: "multiset_of xs = multiset_of ys"
+  shows "multiset_of (sublist' l r xs) = multiset_of (sublist' l r ys)"
+proof -
+  from l_r have xs_def: "xs = (sublist' 0 l xs) @ (sublist' l r xs) @ (sublist' r (List.length xs) xs)" (is "_ = ?xs_long") 
+    by (simp add: sublist'_append)
+  from multiset have length_eq: "List.length xs = List.length ys" by (rule multiset_of_eq_length)
+  with l_r have ys_def: "ys = (sublist' 0 l ys) @ (sublist' l r ys) @ (sublist' r (List.length ys) ys)" (is "_ = ?ys_long") 
+    by (simp add: sublist'_append)
+  from xs_def ys_def multiset have "multiset_of ?xs_long = multiset_of ?ys_long" by simp
+  moreover
+  from left l_r length_eq have "sublist' 0 l xs = sublist' 0 l ys"
+    by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI)
+  moreover
+  from right l_r length_eq have "sublist' r (List.length xs) xs = sublist' r (List.length ys) ys"
+    by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI)
+  moreover
+  ultimately show ?thesis by (simp add: multiset_of_append)
+qed
+
+
+end
--- a/src/HOL/Import/HOL4Compat.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Import/HOL4Compat.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,11 +1,14 @@
 (*  Title:      HOL/Import/HOL4Compat.thy
-    ID:         $Id$
     Author:     Sebastian Skalberg (TU Muenchen)
 *)
 
-theory HOL4Compat imports HOL4Setup Divides Primes Real ContNotDenum
+theory HOL4Compat
+imports HOL4Setup Complex_Main Primes ContNotDenum
 begin
 
+no_notation differentiable (infixl "differentiable" 60)
+no_notation sums (infixr "sums" 80)
+
 lemma EXISTS_UNIQUE_DEF: "(Ex1 P) = (Ex P & (ALL x y. P x & P y --> (x = y)))"
   by auto
 
@@ -22,7 +25,7 @@
 lemmas [hol4rew] = ONE_ONE_rew
 
 lemma bool_case_DEF: "(bool_case x y b) = (if b then x else y)"
-  by simp;
+  by simp
 
 lemma INR_INL_11: "(ALL y x. (Inl x = Inl y) = (x = y)) & (ALL y x. (Inr x = Inr y) = (x = y))"
   by safe
--- a/src/HOL/Int.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Int.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1256,14 +1256,14 @@
 by (simp add: algebra_simps diff_number_of_eq [symmetric])
 
 
+
+
 subsection {* The Set of Integers *}
 
 context ring_1
 begin
 
-definition
-  Ints  :: "'a set"
-where
+definition Ints  :: "'a set" where
   [code del]: "Ints = range of_int"
 
 end
@@ -1854,7 +1854,7 @@
 qed
 
 
-subsection{*Integer Powers*} 
+subsection {* Integer Powers *} 
 
 instantiation int :: recpower
 begin
@@ -1896,6 +1896,161 @@
 
 lemmas zpower_int = int_power [symmetric]
 
+
+subsection {* Further theorems on numerals *}
+
+subsubsection{*Special Simplification for Constants*}
+
+text{*These distributive laws move literals inside sums and differences.*}
+
+lemmas left_distrib_number_of [simp] =
+  left_distrib [of _ _ "number_of v", standard]
+
+lemmas right_distrib_number_of [simp] =
+  right_distrib [of "number_of v", standard]
+
+lemmas left_diff_distrib_number_of [simp] =
+  left_diff_distrib [of _ _ "number_of v", standard]
+
+lemmas right_diff_distrib_number_of [simp] =
+  right_diff_distrib [of "number_of v", standard]
+
+text{*These are actually for fields, like real: but where else to put them?*}
+
+lemmas zero_less_divide_iff_number_of [simp, noatp] =
+  zero_less_divide_iff [of "number_of w", standard]
+
+lemmas divide_less_0_iff_number_of [simp, noatp] =
+  divide_less_0_iff [of "number_of w", standard]
+
+lemmas zero_le_divide_iff_number_of [simp, noatp] =
+  zero_le_divide_iff [of "number_of w", standard]
+
+lemmas divide_le_0_iff_number_of [simp, noatp] =
+  divide_le_0_iff [of "number_of w", standard]
+
+
+text {*Replaces @{text "inverse #nn"} by @{text "1/#nn"}.  It looks
+  strange, but then other simprocs simplify the quotient.*}
+
+lemmas inverse_eq_divide_number_of [simp] =
+  inverse_eq_divide [of "number_of w", standard]
+
+text {*These laws simplify inequalities, moving unary minus from a term
+into the literal.*}
+
+lemmas less_minus_iff_number_of [simp, noatp] =
+  less_minus_iff [of "number_of v", standard]
+
+lemmas le_minus_iff_number_of [simp, noatp] =
+  le_minus_iff [of "number_of v", standard]
+
+lemmas equation_minus_iff_number_of [simp, noatp] =
+  equation_minus_iff [of "number_of v", standard]
+
+lemmas minus_less_iff_number_of [simp, noatp] =
+  minus_less_iff [of _ "number_of v", standard]
+
+lemmas minus_le_iff_number_of [simp, noatp] =
+  minus_le_iff [of _ "number_of v", standard]
+
+lemmas minus_equation_iff_number_of [simp, noatp] =
+  minus_equation_iff [of _ "number_of v", standard]
+
+
+text{*To Simplify Inequalities Where One Side is the Constant 1*}
+
+lemma less_minus_iff_1 [simp,noatp]:
+  fixes b::"'b::{ordered_idom,number_ring}"
+  shows "(1 < - b) = (b < -1)"
+by auto
+
+lemma le_minus_iff_1 [simp,noatp]:
+  fixes b::"'b::{ordered_idom,number_ring}"
+  shows "(1 \<le> - b) = (b \<le> -1)"
+by auto
+
+lemma equation_minus_iff_1 [simp,noatp]:
+  fixes b::"'b::number_ring"
+  shows "(1 = - b) = (b = -1)"
+by (subst equation_minus_iff, auto)
+
+lemma minus_less_iff_1 [simp,noatp]:
+  fixes a::"'b::{ordered_idom,number_ring}"
+  shows "(- a < 1) = (-1 < a)"
+by auto
+
+lemma minus_le_iff_1 [simp,noatp]:
+  fixes a::"'b::{ordered_idom,number_ring}"
+  shows "(- a \<le> 1) = (-1 \<le> a)"
+by auto
+
+lemma minus_equation_iff_1 [simp,noatp]:
+  fixes a::"'b::number_ring"
+  shows "(- a = 1) = (a = -1)"
+by (subst minus_equation_iff, auto)
+
+
+text {*Cancellation of constant factors in comparisons (@{text "<"} and @{text "\<le>"}) *}
+
+lemmas mult_less_cancel_left_number_of [simp, noatp] =
+  mult_less_cancel_left [of "number_of v", standard]
+
+lemmas mult_less_cancel_right_number_of [simp, noatp] =
+  mult_less_cancel_right [of _ "number_of v", standard]
+
+lemmas mult_le_cancel_left_number_of [simp, noatp] =
+  mult_le_cancel_left [of "number_of v", standard]
+
+lemmas mult_le_cancel_right_number_of [simp, noatp] =
+  mult_le_cancel_right [of _ "number_of v", standard]
+
+
+text {*Multiplying out constant divisors in comparisons (@{text "<"}, @{text "\<le>"} and @{text "="}) *}
+
+lemmas le_divide_eq_number_of1 [simp] = le_divide_eq [of _ _ "number_of w", standard]
+lemmas divide_le_eq_number_of1 [simp] = divide_le_eq [of _ "number_of w", standard]
+lemmas less_divide_eq_number_of1 [simp] = less_divide_eq [of _ _ "number_of w", standard]
+lemmas divide_less_eq_number_of1 [simp] = divide_less_eq [of _ "number_of w", standard]
+lemmas eq_divide_eq_number_of1 [simp] = eq_divide_eq [of _ _ "number_of w", standard]
+lemmas divide_eq_eq_number_of1 [simp] = divide_eq_eq [of _ "number_of w", standard]
+
+
+subsubsection{*Optional Simplification Rules Involving Constants*}
+
+text{*Simplify quotients that are compared with a literal constant.*}
+
+lemmas le_divide_eq_number_of = le_divide_eq [of "number_of w", standard]
+lemmas divide_le_eq_number_of = divide_le_eq [of _ _ "number_of w", standard]
+lemmas less_divide_eq_number_of = less_divide_eq [of "number_of w", standard]
+lemmas divide_less_eq_number_of = divide_less_eq [of _ _ "number_of w", standard]
+lemmas eq_divide_eq_number_of = eq_divide_eq [of "number_of w", standard]
+lemmas divide_eq_eq_number_of = divide_eq_eq [of _ _ "number_of w", standard]
+
+
+text{*Not good as automatic simprules because they cause case splits.*}
+lemmas divide_const_simps =
+  le_divide_eq_number_of divide_le_eq_number_of less_divide_eq_number_of
+  divide_less_eq_number_of eq_divide_eq_number_of divide_eq_eq_number_of
+  le_divide_eq_1 divide_le_eq_1 less_divide_eq_1 divide_less_eq_1
+
+text{*Division By @{text "-1"}*}
+
+lemma divide_minus1 [simp]:
+     "x/-1 = -(x::'a::{field,division_by_zero,number_ring})"
+by simp
+
+lemma minus1_divide [simp]:
+     "-1 / (x::'a::{field,division_by_zero,number_ring}) = - (1/x)"
+by (simp add: divide_inverse inverse_minus_eq)
+
+lemma half_gt_zero_iff:
+     "(0 < r/2) = (0 < (r::'a::{ordered_field,division_by_zero,number_ring}))"
+by auto
+
+lemmas half_gt_zero [simp] = half_gt_zero_iff [THEN iffD2, standard]
+
+
 subsection {* Configuration of the code generator *}
 
 code_datatype Pls Min Bit0 Bit1 "number_of \<Colon> int \<Rightarrow> int"
--- a/src/HOL/IntDiv.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/IntDiv.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -8,6 +8,10 @@
 
 theory IntDiv
 imports Int Divides FunDef
+uses
+  "~~/src/Provers/Arith/cancel_numeral_factor.ML"
+  "~~/src/Provers/Arith/extract_common_term.ML"
+  ("Tools/int_factor_simprocs.ML")
 begin
 
 definition divmod_rel :: "int \<Rightarrow> int \<Rightarrow> int \<times> int \<Rightarrow> bool" where
@@ -920,9 +924,10 @@
 next
   assume "a\<noteq>0" and le_a: "0\<le>a"   
   hence a_pos: "1 \<le> a" by arith
-  hence one_less_a2: "1 < 2*a" by arith
+  hence one_less_a2: "1 < 2 * a" by arith
   hence le_2a: "2 * (1 + b mod a) \<le> 2 * a"
-    by (simp add: mult_le_cancel_left add_commute [of 1] add1_zle_eq)
+    unfolding mult_le_cancel_left
+    by (simp add: add1_zle_eq add_commute [of 1])
   with a_pos have "0 \<le> b mod a" by simp
   hence le_addm: "0 \<le> 1 mod (2*a) + 2*(b mod a)"
     by (simp add: mod_pos_pos_trivial one_less_a2)
@@ -1357,6 +1362,11 @@
 qed
 
 
+subsection {* Simproc setup *}
+
+use "Tools/int_factor_simprocs.ML"
+
+
 subsection {* Code generation *}
 
 definition pdivmod :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
--- a/src/HOL/IsaMakefile	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/IsaMakefile	Sat Mar 28 00:13:01 2009 +0100
@@ -204,7 +204,6 @@
 	@$(ISABELLE_TOOL) usedir -b -f plain.ML -g true $(OUT)/Pure HOL-Plain
 
 MAIN_DEPENDENCIES = $(PLAIN_DEPENDENCIES) \
-  Arith_Tools.thy \
   ATP_Linkup.thy \
   Code_Eval.thy \
   Code_Message.thy \
@@ -650,7 +649,11 @@
 $(LOG)/HOL-Imperative_HOL.gz: $(OUT)/HOL Imperative_HOL/Heap.thy \
   Imperative_HOL/Heap_Monad.thy Imperative_HOL/Array.thy \
   Imperative_HOL/Relational.thy \
-  Imperative_HOL/Ref.thy Imperative_HOL/Imperative_HOL.thy
+  Imperative_HOL/Ref.thy Imperative_HOL/Imperative_HOL.thy \
+  Imperative_HOL/Imperative_HOL_ex.thy \
+  Imperative_HOL/ex/Imperative_Quicksort.thy \
+  Imperative_HOL/ex/Subarray.thy \
+  Imperative_HOL/ex/Sublist.thy
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Imperative_HOL
 
 
@@ -837,7 +840,7 @@
   ex/Formal_Power_Series_Examples.thy ex/Fundefs.thy			\
   ex/Groebner_Examples.thy ex/Guess.thy ex/HarmonicSeries.thy		\
   ex/Hebrew.thy ex/Hex_Bin_Examples.thy ex/Higher_Order_Logic.thy	\
-  ex/Hilbert_Classical.thy ex/ImperativeQuicksort.thy			\
+  ex/Hilbert_Classical.thy			\
   ex/Induction_Scheme.thy ex/InductiveInvariant.thy			\
   ex/InductiveInvariant_examples.thy ex/Intuitionistic.thy		\
   ex/Lagrange.thy ex/LocaleTest2.thy ex/MT.thy ex/MergeSort.thy		\
@@ -846,8 +849,8 @@
   ex/Quickcheck_Examples.thy ex/Quickcheck_Generators.thy ex/ROOT.ML	\
   ex/Recdefs.thy ex/Records.thy ex/ReflectionEx.thy			\
   ex/Refute_Examples.thy ex/SAT_Examples.thy ex/SVC_Oracle.thy		\
-  ex/Serbian.thy ex/Sqrt.thy ex/Sqrt_Script.thy ex/Subarray.thy		\
-  ex/Sublist.thy ex/Sudoku.thy ex/Tarski.thy ex/Term_Of_Syntax.thy	\
+  ex/Serbian.thy ex/Sqrt.thy ex/Sqrt_Script.thy \
+  ex/Sudoku.thy ex/Tarski.thy ex/Term_Of_Syntax.thy	\
   ex/Termination.thy ex/Unification.thy ex/document/root.bib		\
   ex/document/root.tex ex/set.thy ex/svc_funcs.ML ex/svc_test.thy \
   ex/Predicate_Compile.thy ex/predicate_compile.ML
--- a/src/HOL/Lattices.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Lattices.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -299,7 +299,7 @@
   by auto
 qed (auto simp add: min_def max_def not_le less_imp_le)
 
-interpretation min_max!: distrib_lattice "op \<le> :: 'a::linorder \<Rightarrow> 'a \<Rightarrow> bool" "op <" min max
+interpretation min_max: distrib_lattice "op \<le> :: 'a::linorder \<Rightarrow> 'a \<Rightarrow> bool" "op <" min max
   by (rule distrib_lattice_min_max)
 
 lemma inf_min: "inf = (min \<Colon> 'a\<Colon>{lower_semilattice, linorder} \<Rightarrow> 'a \<Rightarrow> 'a)"
--- a/src/HOL/Library/Abstract_Rat.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Abstract_Rat.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/Abstract_Rat.thy
-    ID:         $Id$
     Author:     Amine Chaieb
 *)
 
 header {* Abstract rational numbers *}
 
 theory Abstract_Rat
-imports Plain GCD
+imports GCD Main
 begin
 
 types Num = "int \<times> int"
--- a/src/HOL/Library/AssocList.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/AssocList.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/AssocList.thy
-    ID:         $Id$
     Author:     Norbert Schirmer, Tobias Nipkow, Martin Wildmoser
 *)
 
 header {* Map operations implemented on association lists*}
 
 theory AssocList 
-imports Plain "~~/src/HOL/Map"
+imports Map Main
 begin
 
 text {*
--- a/src/HOL/Library/Binomial.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Binomial.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -6,7 +6,7 @@
 header {* Binomial Coefficients *}
 
 theory Binomial
-imports Fact Plain "~~/src/HOL/SetInterval" Presburger 
+imports Fact SetInterval Presburger Main
 begin
 
 text {* This development is based on the work of Andy Gordon and
--- a/src/HOL/Library/Boolean_Algebra.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Boolean_Algebra.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Boolean Algebras *}
 
 theory Boolean_Algebra
-imports Plain
+imports Main
 begin
 
 locale boolean =
--- a/src/HOL/Library/Char_nat.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Char_nat.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/Char_nat.thy
-    ID:         $Id$
     Author:     Norbert Voelker, Florian Haftmann
 *)
 
 header {* Mapping between characters and natural numbers *}
 
 theory Char_nat
-imports Plain "~~/src/HOL/List"
+imports List Main
 begin
 
 text {* Conversions between nibbles and natural numbers in [0..15]. *}
--- a/src/HOL/Library/Char_ord.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Char_ord.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/Char_ord.thy
-    ID:         $Id$
     Author:     Norbert Voelker, Florian Haftmann
 *)
 
 header {* Order on characters *}
 
 theory Char_ord
-imports Plain Product_ord Char_nat
+imports Product_ord Char_nat Main
 begin
 
 instantiation nibble :: linorder
--- a/src/HOL/Library/Code_Char.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Code_Char.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Code generation of pretty characters (and strings) *}
 
 theory Code_Char
-imports Plain "~~/src/HOL/List" "~~/src/HOL/Code_Eval"
+imports List Code_Eval Main
 begin
 
 code_type char
--- a/src/HOL/Library/Code_Char_chr.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Code_Char_chr.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/Code_Char_chr.thy
-    ID:         $Id$
     Author:     Florian Haftmann
 *)
 
 header {* Code generation of pretty characters with character codes *}
 
 theory Code_Char_chr
-imports Plain Char_nat Code_Char Code_Integer
+imports Char_nat Code_Char Code_Integer Main
 begin
 
 definition
--- a/src/HOL/Library/Code_Index.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Code_Index.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -3,7 +3,7 @@
 header {* Type of indices *}
 
 theory Code_Index
-imports Plain "~~/src/HOL/Code_Eval" "~~/src/HOL/Presburger"
+imports Main
 begin
 
 text {*
--- a/src/HOL/Library/Code_Integer.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Code_Integer.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Pretty integer literals for code generation *}
 
 theory Code_Integer
-imports Plain "~~/src/HOL/Code_Eval" "~~/src/HOL/Presburger"
+imports Main
 begin
 
 text {*
--- a/src/HOL/Library/Coinductive_List.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Coinductive_List.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/Coinductive_Lists.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson and Makarius
 *)
 
 header {* Potentially infinite lists as greatest fixed-point *}
 
 theory Coinductive_List
-imports Plain "~~/src/HOL/List"
+imports List Main
 begin
 
 subsection {* List constructors over the datatype universe *}
--- a/src/HOL/Library/Commutative_Ring.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Commutative_Ring.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -6,7 +6,7 @@
 header {* Proving equalities in commutative rings *}
 
 theory Commutative_Ring
-imports Plain "~~/src/HOL/List" "~~/src/HOL/Parity"
+imports List Parity Main
 uses ("comm_ring.ML")
 begin
 
--- a/src/HOL/Library/ContNotDenum.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/ContNotDenum.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Non-denumerability of the Continuum. *}
 
 theory ContNotDenum
-imports RComplete Hilbert_Choice
+imports Complex_Main
 begin
 
 subsection {* Abstract *}
--- a/src/HOL/Library/Continuity.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Continuity.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/Continuity.thy
-    ID:         $Id$
     Author:     David von Oheimb, TU Muenchen
 *)
 
 header {* Continuity and iterations (of set transformers) *}
 
 theory Continuity
-imports Plain "~~/src/HOL/Relation_Power"
+imports Relation_Power Main
 begin
 
 subsection {* Continuity for complete lattices *}
--- a/src/HOL/Library/Countable.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Countable.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -6,11 +6,11 @@
 
 theory Countable
 imports
-  Plain
   "~~/src/HOL/List"
   "~~/src/HOL/Hilbert_Choice"
   "~~/src/HOL/Nat_Int_Bij"
   "~~/src/HOL/Rational"
+  Main
 begin
 
 subsection {* The class of countable types *}
--- a/src/HOL/Library/Determinants.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Determinants.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Traces, Determinant of square matrices and some properties *}
 
 theory Determinants
-  imports Euclidean_Space Permutations
+imports Euclidean_Space Permutations
 begin
 
 subsection{* First some facts about products*}
@@ -68,22 +68,22 @@
 subsection{* Trace *}
 
 definition trace :: "'a::semiring_1^'n^'n \<Rightarrow> 'a" where
-  "trace A = setsum (\<lambda>i. ((A$i)$i)) {1..dimindex(UNIV::'n set)}"
+  "trace A = setsum (\<lambda>i. ((A$i)$i)) (UNIV::'n set)"
 
 lemma trace_0: "trace(mat 0) = 0"
-  by (simp add: trace_def mat_def Cart_lambda_beta setsum_0)
+  by (simp add: trace_def mat_def)
 
-lemma trace_I: "trace(mat 1 :: 'a::semiring_1^'n^'n) = of_nat(dimindex(UNIV::'n set))"
-  by (simp add: trace_def mat_def Cart_lambda_beta)
+lemma trace_I: "trace(mat 1 :: 'a::semiring_1^'n^'n) = of_nat(CARD('n))"
+  by (simp add: trace_def mat_def)
 
 lemma trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B"
-  by (simp add: trace_def setsum_addf Cart_lambda_beta vector_component)
+  by (simp add: trace_def setsum_addf)
 
 lemma trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B"
-  by (simp add: trace_def setsum_subtractf Cart_lambda_beta vector_component)
+  by (simp add: trace_def setsum_subtractf)
 
 lemma trace_mul_sym:"trace ((A::'a::comm_semiring_1^'n^'n) ** B) = trace (B**A)"
-  apply (simp add: trace_def matrix_matrix_mult_def Cart_lambda_beta)
+  apply (simp add: trace_def matrix_matrix_mult_def)
   apply (subst setsum_commute)
   by (simp add: mult_commute)
 
@@ -92,18 +92,12 @@
 (* ------------------------------------------------------------------------- *)
 
 definition det:: "'a::comm_ring_1^'n^'n \<Rightarrow> 'a" where
-  "det A = setsum (\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) {1 .. dimindex(UNIV :: 'n set)}) {p. p permutes {1 .. dimindex(UNIV :: 'n set)}}"
+  "det A = setsum (\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)) {p. p permutes (UNIV :: 'n set)}"
 
 (* ------------------------------------------------------------------------- *)
 (* A few general lemmas we need below.                                       *)
 (* ------------------------------------------------------------------------- *)
 
-lemma Cart_lambda_beta_perm: assumes p: "p permutes {1..dimindex(UNIV::'n set)}"
-  and i: "i \<in> {1..dimindex(UNIV::'n set)}"
-  shows "Cart_nth (Cart_lambda g ::'a^'n) (p i) = g(p i)"
-  using permutes_in_image[OF p] i
-  by (simp add:  Cart_lambda_beta permutes_in_image[OF p])
-
 lemma setprod_permute:
   assumes p: "p permutes S"
   shows "setprod f S = setprod (f o p) S"
@@ -127,11 +121,11 @@
 (* Basic determinant properties.                                             *)
 (* ------------------------------------------------------------------------- *)
 
-lemma det_transp: "det (transp A) = det (A::'a::comm_ring_1 ^'n^'n)"
+lemma det_transp: "det (transp A) = det (A::'a::comm_ring_1 ^'n^'n::finite)"
 proof-
   let ?di = "\<lambda>A i j. A$i$j"
-  let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
-  have fU: "finite ?U" by blast
+  let ?U = "(UNIV :: 'n set)"
+  have fU: "finite ?U" by simp
   {fix p assume p: "p \<in> {p. p permutes ?U}"
     from p have pU: "p permutes ?U" by blast
     have sth: "sign (inv p) = sign p"
@@ -147,7 +141,7 @@
       {fix i assume i: "i \<in> ?U"
 	from i permutes_inv_o[OF pU] permutes_in_image[OF pU]
 	have "((\<lambda>i. ?di (transp A) i (inv p i)) o p) i = ?di A i (p i)"
-	  unfolding transp_def by (simp add: Cart_lambda_beta expand_fun_eq)}
+	  unfolding transp_def by (simp add: expand_fun_eq)}
       then show "setprod ((\<lambda>i. ?di (transp A) i (inv p i)) o p) ?U = setprod (\<lambda>i. ?di A i (p i)) ?U" by (auto intro: setprod_cong)
     qed
     finally have "of_int (sign (inv p)) * (setprod (\<lambda>i. ?di (transp A) i (inv p i)) ?U) = of_int (sign p) * (setprod (\<lambda>i. ?di A i (p i)) ?U)" using sth
@@ -157,22 +151,21 @@
 qed
 
 lemma det_lowerdiagonal:
-  fixes A :: "'a::comm_ring_1^'n^'n"
-  assumes ld: "\<And>i j. i \<in> {1 .. dimindex (UNIV:: 'n set)} \<Longrightarrow> j \<in> {1 .. dimindex(UNIV:: 'n set)} \<Longrightarrow> i < j \<Longrightarrow> A$i$j = 0"
-  shows "det A = setprod (\<lambda>i. A$i$i) {1..dimindex(UNIV:: 'n set)}"
+  fixes A :: "'a::comm_ring_1^'n^'n::{finite,wellorder}"
+  assumes ld: "\<And>i j. i < j \<Longrightarrow> A$i$j = 0"
+  shows "det A = setprod (\<lambda>i. A$i$i) (UNIV:: 'n set)"
 proof-
-  let ?U = "{1..dimindex(UNIV:: 'n set)}"
+  let ?U = "UNIV:: 'n set"
   let ?PU = "{p. p permutes ?U}"
-  let ?pp = "\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) {1 .. dimindex(UNIV :: 'n set)}"
-  have fU: "finite ?U" by blast
+  let ?pp = "\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)"
+  have fU: "finite ?U" by simp
   from finite_permutations[OF fU] have fPU: "finite ?PU" .
   have id0: "{id} \<subseteq> ?PU" by (auto simp add: permutes_id)
   {fix p assume p: "p \<in> ?PU -{id}"
     from p have pU: "p permutes ?U" and pid: "p \<noteq> id" by blast+
     from permutes_natset_le[OF pU] pid obtain i where
-      i: "i \<in> ?U" "p i > i" by (metis not_le)
-    from permutes_in_image[OF pU] i(1) have piU: "p i \<in> ?U" by blast
-    from ld[OF i(1) piU i(2)] i(1) have ex:"\<exists>i \<in> ?U. A$i$p i = 0" by blast
+      i: "p i > i" by (metis not_le)
+    from ld[OF i] have ex:"\<exists>i \<in> ?U. A$i$p i = 0" by blast
     from setprod_zero[OF fU ex] have "?pp p = 0" by simp}
   then have p0: "\<forall>p \<in> ?PU -{id}. ?pp p = 0"  by blast
   from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
@@ -180,99 +173,97 @@
 qed
 
 lemma det_upperdiagonal:
-  fixes A :: "'a::comm_ring_1^'n^'n"
-  assumes ld: "\<And>i j. i \<in> {1 .. dimindex (UNIV:: 'n set)} \<Longrightarrow> j \<in> {1 .. dimindex(UNIV:: 'n set)} \<Longrightarrow> i > j \<Longrightarrow> A$i$j = 0"
-  shows "det A = setprod (\<lambda>i. A$i$i) {1..dimindex(UNIV:: 'n set)}"
+  fixes A :: "'a::comm_ring_1^'n^'n::{finite,wellorder}"
+  assumes ld: "\<And>i j. i > j \<Longrightarrow> A$i$j = 0"
+  shows "det A = setprod (\<lambda>i. A$i$i) (UNIV:: 'n set)"
 proof-
-  let ?U = "{1..dimindex(UNIV:: 'n set)}"
+  let ?U = "UNIV:: 'n set"
   let ?PU = "{p. p permutes ?U}"
-  let ?pp = "(\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) {1 .. dimindex(UNIV :: 'n set)})"
-  have fU: "finite ?U" by blast
+  let ?pp = "(\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set))"
+  have fU: "finite ?U" by simp
   from finite_permutations[OF fU] have fPU: "finite ?PU" .
   have id0: "{id} \<subseteq> ?PU" by (auto simp add: permutes_id)
   {fix p assume p: "p \<in> ?PU -{id}"
     from p have pU: "p permutes ?U" and pid: "p \<noteq> id" by blast+
     from permutes_natset_ge[OF pU] pid obtain i where
-      i: "i \<in> ?U" "p i < i" by (metis not_le)
-    from permutes_in_image[OF pU] i(1) have piU: "p i \<in> ?U" by blast
-    from ld[OF i(1) piU i(2)] i(1) have ex:"\<exists>i \<in> ?U. A$i$p i = 0" by blast
+      i: "p i < i" by (metis not_le)
+    from ld[OF i] have ex:"\<exists>i \<in> ?U. A$i$p i = 0" by blast
     from setprod_zero[OF fU ex] have "?pp p = 0" by simp}
   then have p0: "\<forall>p \<in> ?PU -{id}. ?pp p = 0"  by blast
   from   setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
     unfolding det_def by (simp add: sign_id)
 qed
 
-lemma det_I: "det (mat 1 :: 'a::comm_ring_1^'n^'n) = 1"
+lemma det_diagonal:
+  fixes A :: "'a::comm_ring_1^'n^'n::finite"
+  assumes ld: "\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0"
+  shows "det A = setprod (\<lambda>i. A$i$i) (UNIV::'n set)"
+proof-
+  let ?U = "UNIV:: 'n set"
+  let ?PU = "{p. p permutes ?U}"
+  let ?pp = "\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)"
+  have fU: "finite ?U" by simp
+  from finite_permutations[OF fU] have fPU: "finite ?PU" .
+  have id0: "{id} \<subseteq> ?PU" by (auto simp add: permutes_id)
+  {fix p assume p: "p \<in> ?PU - {id}"
+    then have "p \<noteq> id" by simp
+    then obtain i where i: "p i \<noteq> i" unfolding expand_fun_eq by auto
+    from ld [OF i [symmetric]] have ex:"\<exists>i \<in> ?U. A$i$p i = 0" by blast
+    from setprod_zero [OF fU ex] have "?pp p = 0" by simp}
+  then have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"  by blast
+  from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
+    unfolding det_def by (simp add: sign_id)
+qed
+
+lemma det_I: "det (mat 1 :: 'a::comm_ring_1^'n^'n::finite) = 1"
 proof-
   let ?A = "mat 1 :: 'a::comm_ring_1^'n^'n"
-  let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
+  let ?U = "UNIV :: 'n set"
   let ?f = "\<lambda>i j. ?A$i$j"
   {fix i assume i: "i \<in> ?U"
     have "?f i i = 1" using i by (vector mat_def)}
   hence th: "setprod (\<lambda>i. ?f i i) ?U = setprod (\<lambda>x. 1) ?U"
     by (auto intro: setprod_cong)
-  {fix i j assume i: "i \<in> ?U" and j: "j \<in> ?U" and ij: "i < j"
+  {fix i j assume i: "i \<in> ?U" and j: "j \<in> ?U" and ij: "i \<noteq> j"
     have "?f i j = 0" using i j ij by (vector mat_def) }
-  then have "det ?A = setprod (\<lambda>i. ?f i i) ?U" using det_lowerdiagonal
+  then have "det ?A = setprod (\<lambda>i. ?f i i) ?U" using det_diagonal
     by blast
   also have "\<dots> = 1" unfolding th setprod_1 ..
   finally show ?thesis .
 qed
 
-lemma det_0: "det (mat 0 :: 'a::comm_ring_1^'n^'n) = 0"
-proof-
-  let ?A = "mat 0 :: 'a::comm_ring_1^'n^'n"
-  let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
-  let ?f = "\<lambda>i j. ?A$i$j"
-  have th:"setprod (\<lambda>i. ?f i i) ?U = 0"
-    apply (rule setprod_zero)
-    apply simp
-    apply (rule bexI[where x=1])
-    using dimindex_ge_1[of "UNIV :: 'n set"]
-    by (simp_all add: mat_def Cart_lambda_beta)
-  {fix i j assume i: "i \<in> ?U" and j: "j \<in> ?U" and ij: "i < j"
-    have "?f i j = 0" using i j ij by (vector mat_def) }
-  then have "det ?A = setprod (\<lambda>i. ?f i i) ?U" using det_lowerdiagonal
-    by blast
-  also have "\<dots> = 0" unfolding th  ..
-  finally show ?thesis .
-qed
+lemma det_0: "det (mat 0 :: 'a::comm_ring_1^'n^'n::finite) = 0"
+  by (simp add: det_def setprod_zero)
 
 lemma det_permute_rows:
-  fixes A :: "'a::comm_ring_1^'n^'n"
-  assumes p: "p permutes {1 .. dimindex (UNIV :: 'n set)}"
+  fixes A :: "'a::comm_ring_1^'n^'n::finite"
+  assumes p: "p permutes (UNIV :: 'n::finite set)"
   shows "det(\<chi> i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A"
-  apply (simp add: det_def setsum_right_distrib mult_assoc[symmetric] del: One_nat_def)
+  apply (simp add: det_def setsum_right_distrib mult_assoc[symmetric])
   apply (subst sum_permutations_compose_right[OF p])
 proof(rule setsum_cong2)
-  let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
+  let ?U = "UNIV :: 'n set"
   let ?PU = "{p. p permutes ?U}"
-  let ?Ap = "(\<chi> i. A$p i :: 'a^'n^'n)"
   fix q assume qPU: "q \<in> ?PU"
-  have fU: "finite ?U" by blast
+  have fU: "finite ?U" by simp
   from qPU have q: "q permutes ?U" by blast
   from p q have pp: "permutation p" and qp: "permutation q"
     by (metis fU permutation_permutes)+
   from permutes_inv[OF p] have ip: "inv p permutes ?U" .
-    {fix i assume i: "i \<in> ?U"
-      from Cart_lambda_beta[rule_format, OF i, of "\<lambda>i. A$ p i"]
-      have "?Ap$i$ (q o p) i = A $ p i $ (q o p) i " by simp}
-    hence "setprod (\<lambda>i. ?Ap$i$ (q o p) i) ?U = setprod (\<lambda>i. A$p i$(q o p) i) ?U"
-      by (auto intro: setprod_cong)
-    also have "\<dots> = setprod ((\<lambda>i. A$p i$(q o p) i) o inv p) ?U"
+    have "setprod (\<lambda>i. A$p i$ (q o p) i) ?U = setprod ((\<lambda>i. A$p i$(q o p) i) o inv p) ?U"
       by (simp only: setprod_permute[OF ip, symmetric])
     also have "\<dots> = setprod (\<lambda>i. A $ (p o inv p) i $ (q o (p o inv p)) i) ?U"
       by (simp only: o_def)
     also have "\<dots> = setprod (\<lambda>i. A$i$q i) ?U" by (simp only: o_def permutes_inverses[OF p])
-    finally   have thp: "setprod (\<lambda>i. ?Ap$i$ (q o p) i) ?U = setprod (\<lambda>i. A$i$q i) ?U"
+    finally   have thp: "setprod (\<lambda>i. A$p i$ (q o p) i) ?U = setprod (\<lambda>i. A$i$q i) ?U"
       by blast
-  show "of_int (sign (q o p)) * setprod (\<lambda>i. ?Ap$i$ (q o p) i) ?U = of_int (sign p) * of_int (sign q) * setprod (\<lambda>i. A$i$q i) ?U"
+  show "of_int (sign (q o p)) * setprod (\<lambda>i. A$ p i$ (q o p) i) ?U = of_int (sign p) * of_int (sign q) * setprod (\<lambda>i. A$i$q i) ?U"
     by (simp only: thp sign_compose[OF qp pp] mult_commute of_int_mult)
 qed
 
 lemma det_permute_columns:
-  fixes A :: "'a::comm_ring_1^'n^'n"
-  assumes p: "p permutes {1 .. dimindex (UNIV :: 'n set)}"
+  fixes A :: "'a::comm_ring_1^'n^'n::finite"
+  assumes p: "p permutes (UNIV :: 'n set)"
   shows "det(\<chi> i j. A$i$ p j :: 'a^'n^'n) = of_int (sign p) * det A"
 proof-
   let ?Ap = "\<chi> i j. A$i$ p j :: 'a^'n^'n"
@@ -281,15 +272,13 @@
     unfolding det_permute_rows[OF p, of ?At] det_transp ..
   moreover
   have "?Ap = transp (\<chi> i. transp A $ p i)"
-    by (simp add: transp_def Cart_eq Cart_lambda_beta Cart_lambda_beta_perm[OF p])
+    by (simp add: transp_def Cart_eq)
   ultimately show ?thesis by simp
 qed
 
 lemma det_identical_rows:
-  fixes A :: "'a::ordered_idom^'n^'n"
-  assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}"
-  and j: "j\<in>{1 .. dimindex (UNIV :: 'n set)}"
-  and ij: "i \<noteq> j"
+  fixes A :: "'a::ordered_idom^'n^'n::finite"
+  assumes ij: "i \<noteq> j"
   and r: "row i A = row j A"
   shows	"det A = 0"
 proof-
@@ -298,94 +287,59 @@
   have th1: "of_int (-1) = - 1" by (metis of_int_1 of_int_minus number_of_Min)
   let ?p = "Fun.swap i j id"
   let ?A = "\<chi> i. A $ ?p i"
-  from r have "A = ?A" by (simp add: Cart_eq Cart_lambda_beta Cart_lambda_beta_perm[OF permutes_swap_id[OF i j]] row_def swap_def)
+  from r have "A = ?A" by (simp add: Cart_eq row_def swap_def)
   hence "det A = det ?A" by simp
   moreover have "det A = - det ?A"
-    by (simp add: det_permute_rows[OF permutes_swap_id[OF i j]] sign_swap_id ij th1)
+    by (simp add: det_permute_rows[OF permutes_swap_id] sign_swap_id ij th1)
   ultimately show "det A = 0" by (metis tha)
 qed
 
 lemma det_identical_columns:
-  fixes A :: "'a::ordered_idom^'n^'n"
-  assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}"
-  and j: "j\<in>{1 .. dimindex (UNIV :: 'n set)}"
-  and ij: "i \<noteq> j"
+  fixes A :: "'a::ordered_idom^'n^'n::finite"
+  assumes ij: "i \<noteq> j"
   and r: "column i A = column j A"
   shows	"det A = 0"
 apply (subst det_transp[symmetric])
-apply (rule det_identical_rows[OF i j ij])
-by (metis row_transp i j r)
+apply (rule det_identical_rows[OF ij])
+by (metis row_transp r)
 
 lemma det_zero_row:
-  fixes A :: "'a::{idom, ring_char_0}^'n^'n"
-  assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}"
-  and r: "row i A = 0"
+  fixes A :: "'a::{idom, ring_char_0}^'n^'n::finite"
+  assumes r: "row i A = 0"
   shows "det A = 0"
-using i r
-apply (simp add: row_def det_def Cart_lambda_beta Cart_eq vector_component del: One_nat_def)
+using r
+apply (simp add: row_def det_def Cart_eq)
 apply (rule setsum_0')
-apply (clarsimp simp add: sign_nz simp del: One_nat_def)
+apply (clarsimp simp add: sign_nz)
 apply (rule setprod_zero)
 apply simp
-apply (rule bexI[where x=i])
-apply (erule_tac x="a i" in ballE)
-apply (subgoal_tac "(0\<Colon>'a ^ 'n) $ a i = 0")
-apply simp
-apply (rule zero_index)
-apply (drule permutes_in_image[of _ _ i])
-apply simp
-apply (drule permutes_in_image[of _ _ i])
-apply simp
-apply simp
+apply auto
 done
 
 lemma det_zero_column:
-  fixes A :: "'a::{idom,ring_char_0}^'n^'n"
-  assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}"
-  and r: "column i A = 0"
+  fixes A :: "'a::{idom,ring_char_0}^'n^'n::finite"
+  assumes r: "column i A = 0"
   shows "det A = 0"
   apply (subst det_transp[symmetric])
-  apply (rule det_zero_row[OF i])
-  by (metis row_transp r i)
-
-lemma setsum_lambda_beta[simp]: "setsum (\<lambda>i. ((\<chi> i. g i) :: 'a::{comm_monoid_add}^'n) $ i ) {1 .. dimindex (UNIV :: 'n set)} = setsum g {1 .. dimindex (UNIV :: 'n set)}"
-  by (simp add: Cart_lambda_beta)
-
-lemma setprod_lambda_beta[simp]: "setprod (\<lambda>i. ((\<chi> i. g i) :: 'a::{comm_monoid_mult}^'n) $ i ) {1 .. dimindex (UNIV :: 'n set)} = setprod g {1 .. dimindex (UNIV :: 'n set)}"
-  apply (rule setprod_cong)
-  apply simp
-  apply (simp add: Cart_lambda_beta')
-  done
-
-lemma setprod_lambda_beta2[simp]: "setprod (\<lambda>i. ((\<chi> i. g i) :: 'a::{comm_monoid_mult}^'n^'n) $ i$ f i ) {1 .. dimindex (UNIV :: 'n set)} = setprod (\<lambda>i. g i $ f i) {1 .. dimindex (UNIV :: 'n set)}"
-proof(rule setprod_cong[OF refl])
-  let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
-  fix i assume i: "i \<in> ?U"
-  from Cart_lambda_beta'[OF i, of g] have
-    "((\<chi> i. g i) :: 'a^'n^'n) $ i = g i" .
-  hence "((\<chi> i. g i) :: 'a^'n^'n) $ i $ f i = g i $ f i" by simp
-  then
-  show "((\<chi> i. g i):: 'a^'n^'n) $ i $ f i = g i $ f i"   .
-qed
+  apply (rule det_zero_row [of i])
+  by (metis row_transp r)
 
 lemma det_row_add:
-  assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}"
+  fixes a b c :: "'n::finite \<Rightarrow> _ ^ 'n"
   shows "det((\<chi> i. if i = k then a i + b i else c i)::'a::comm_ring_1^'n^'n) =
              det((\<chi> i. if i = k then a i else c i)::'a::comm_ring_1^'n^'n) +
              det((\<chi> i. if i = k then b i else c i)::'a::comm_ring_1^'n^'n)"
-unfolding det_def setprod_lambda_beta2 setsum_addf[symmetric]
+unfolding det_def Cart_lambda_beta setsum_addf[symmetric]
 proof (rule setsum_cong2)
-  let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
+  let ?U = "UNIV :: 'n set"
   let ?pU = "{p. p permutes ?U}"
-  let ?f = "(\<lambda>i. if i = k then a i + b i else c i)::nat \<Rightarrow> 'a::comm_ring_1^'n"
-  let ?g = "(\<lambda> i. if i = k then a i else c i)::nat \<Rightarrow> 'a::comm_ring_1^'n"
-  let ?h = "(\<lambda> i. if i = k then b i else c i)::nat \<Rightarrow> 'a::comm_ring_1^'n"
+  let ?f = "(\<lambda>i. if i = k then a i + b i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
+  let ?g = "(\<lambda> i. if i = k then a i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
+  let ?h = "(\<lambda> i. if i = k then b i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
   fix p assume p: "p \<in> ?pU"
   let ?Uk = "?U - {k}"
   from p have pU: "p permutes ?U" by blast
-  from k have pkU: "p k \<in> ?U" by (simp only: permutes_in_image[OF pU])
-  note pin[simp] = permutes_in_image[OF pU]
-  have kU: "?U = insert k ?Uk" using k by blast
+  have kU: "?U = insert k ?Uk" by blast
   {fix j assume j: "j \<in> ?Uk"
     from j have "?f j $ p j = ?g j $ p j" and "?f j $ p j= ?h j $ p j"
       by simp_all}
@@ -394,14 +348,14 @@
     apply -
     apply (rule setprod_cong, simp_all)+
     done
-  have th3: "finite ?Uk" "k \<notin> ?Uk" using k by auto
+  have th3: "finite ?Uk" "k \<notin> ?Uk" by auto
   have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?f i $ p i) (insert k ?Uk)"
     unfolding kU[symmetric] ..
   also have "\<dots> = ?f k $ p k  * setprod (\<lambda>i. ?f i $ p i) ?Uk"
     apply (rule setprod_insert)
     apply simp
-    using k by blast
-  also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk)" using pkU by (simp add: ring_simps vector_component)
+    by blast
+  also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk)" by (simp add: ring_simps)
   also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?h i $ p i) ?Uk)" by (metis th1 th2)
   also have "\<dots> = setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk) + setprod (\<lambda>i. ?h i $ p i) (insert k ?Uk)"
     unfolding  setprod_insert[OF th3] by simp
@@ -411,38 +365,36 @@
 qed
 
 lemma det_row_mul:
-  assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}"
+  fixes a b :: "'n::finite \<Rightarrow> _ ^ 'n"
   shows "det((\<chi> i. if i = k then c *s a i else b i)::'a::comm_ring_1^'n^'n) =
              c* det((\<chi> i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)"
 
-unfolding det_def setprod_lambda_beta2 setsum_right_distrib
+unfolding det_def Cart_lambda_beta setsum_right_distrib
 proof (rule setsum_cong2)
-  let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
+  let ?U = "UNIV :: 'n set"
   let ?pU = "{p. p permutes ?U}"
-  let ?f = "(\<lambda>i. if i = k then c*s a i else b i)::nat \<Rightarrow> 'a::comm_ring_1^'n"
-  let ?g = "(\<lambda> i. if i = k then a i else b i)::nat \<Rightarrow> 'a::comm_ring_1^'n"
+  let ?f = "(\<lambda>i. if i = k then c*s a i else b i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
+  let ?g = "(\<lambda> i. if i = k then a i else b i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
   fix p assume p: "p \<in> ?pU"
   let ?Uk = "?U - {k}"
   from p have pU: "p permutes ?U" by blast
-  from k have pkU: "p k \<in> ?U" by (simp only: permutes_in_image[OF pU])
-  note pin[simp] = permutes_in_image[OF pU]
-  have kU: "?U = insert k ?Uk" using k by blast
+  have kU: "?U = insert k ?Uk" by blast
   {fix j assume j: "j \<in> ?Uk"
     from j have "?f j $ p j = ?g j $ p j" by simp}
   then have th1: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?g i $ p i) ?Uk"
     apply -
     apply (rule setprod_cong, simp_all)
     done
-  have th3: "finite ?Uk" "k \<notin> ?Uk" using k by auto
+  have th3: "finite ?Uk" "k \<notin> ?Uk" by auto
   have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?f i $ p i) (insert k ?Uk)"
     unfolding kU[symmetric] ..
   also have "\<dots> = ?f k $ p k  * setprod (\<lambda>i. ?f i $ p i) ?Uk"
     apply (rule setprod_insert)
     apply simp
-    using k by blast
-  also have "\<dots> = (c*s a k) $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" using pkU by (simp add: ring_simps vector_component)
+    by blast
+  also have "\<dots> = (c*s a k) $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" by (simp add: ring_simps)
   also have "\<dots> = c* (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk)"
-    unfolding th1 using pkU by (simp add: vector_component mult_ac)
+    unfolding th1 by (simp add: mult_ac)
   also have "\<dots> = c* (setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk))"
     unfolding  setprod_insert[OF th3] by simp
   finally have "setprod (\<lambda>i. ?f i $ p i) ?U = c* (setprod (\<lambda>i. ?g i $ p i) ?U)" unfolding kU[symmetric] .
@@ -451,36 +403,33 @@
 qed
 
 lemma det_row_0:
-  assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}"
+  fixes b :: "'n::finite \<Rightarrow> _ ^ 'n"
   shows "det((\<chi> i. if i = k then 0 else b i)::'a::comm_ring_1^'n^'n) = 0"
-using det_row_mul[OF k, of 0 "\<lambda>i. 1" b]
+using det_row_mul[of k 0 "\<lambda>i. 1" b]
 apply (simp)
   unfolding vector_smult_lzero .
 
 lemma det_row_operation:
-  fixes A :: "'a::ordered_idom^'n^'n"
-  assumes i: "i \<in> {1 .. dimindex(UNIV :: 'n set)}"
-  and j: "j \<in> {1 .. dimindex(UNIV :: 'n set)}"
-  and ij: "i \<noteq> j"
+  fixes A :: "'a::ordered_idom^'n^'n::finite"
+  assumes ij: "i \<noteq> j"
   shows "det (\<chi> k. if k = i then row i A + c *s row j A else row k A) = det A"
 proof-
   let ?Z = "(\<chi> k. if k = i then row j A else row k A) :: 'a ^'n^'n"
-  have th: "row i ?Z = row j ?Z" using i j by (vector row_def)
+  have th: "row i ?Z = row j ?Z" by (vector row_def)
   have th2: "((\<chi> k. if k = i then row i A else row k A) :: 'a^'n^'n) = A"
-    using i j by (vector row_def)
+    by (vector row_def)
   show ?thesis
-    unfolding det_row_add [OF i] det_row_mul[OF i] det_identical_rows[OF i j ij th] th2
+    unfolding det_row_add [of i] det_row_mul[of i] det_identical_rows[OF ij th] th2
     by simp
 qed
 
 lemma det_row_span:
-  fixes A :: "'a:: ordered_idom^'n^'n"
-  assumes i: "i \<in> {1 .. dimindex(UNIV :: 'n set)}"
-  and x: "x \<in> span {row j A |j. j\<in> {1 .. dimindex(UNIV :: 'n set)} \<and> j\<noteq> i}"
+  fixes A :: "'a:: ordered_idom^'n^'n::finite"
+  assumes x: "x \<in> span {row j A |j. j \<noteq> i}"
   shows "det (\<chi> k. if k = i then row i A + x else row k A) = det A"
 proof-
-  let ?U = "{1 .. dimindex(UNIV :: 'n set)}"
-  let ?S = "{row j A |j. j\<in> ?U \<and> j\<noteq> i}"
+  let ?U = "UNIV :: 'n set"
+  let ?S = "{row j A |j. j \<noteq> i}"
   let ?d = "\<lambda>x. det (\<chi> k. if k = i then x else row k A)"
   let ?P = "\<lambda>x. ?d (row i A + x) = det A"
   {fix k
@@ -489,17 +438,17 @@
   then have P0: "?P 0"
     apply -
     apply (rule cong[of det, OF refl])
-    using i by (vector row_def)
+    by (vector row_def)
   moreover
   {fix c z y assume zS: "z \<in> ?S" and Py: "?P y"
-    from zS obtain j where j: "z = row j A" "j \<in> ?U" "i \<noteq> j" by blast
+    from zS obtain j where j: "z = row j A" "i \<noteq> j" by blast
     let ?w = "row i A + y"
     have th0: "row i A + (c*s z + y) = ?w + c*s z" by vector
     have thz: "?d z = 0"
-      apply (rule det_identical_rows[OF i j(2,3)])
-      using i j by (vector row_def)
+      apply (rule det_identical_rows[OF j(2)])
+      using j by (vector row_def)
     have "?d (row i A + (c*s z + y)) = ?d (?w + c*s z)" unfolding th0 ..
-    then have "?P (c*s z + y)" unfolding thz Py det_row_mul[OF i] det_row_add[OF i]
+    then have "?P (c*s z + y)" unfolding thz Py det_row_mul[of i] det_row_add[of i]
       by simp }
 
   ultimately show ?thesis
@@ -516,48 +465,47 @@
 (* ------------------------------------------------------------------------- *)
 
 lemma det_dependent_rows:
-  fixes A:: "'a::ordered_idom^'n^'n"
+  fixes A:: "'a::ordered_idom^'n^'n::finite"
   assumes d: "dependent (rows A)"
   shows "det A = 0"
 proof-
-  let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
-  from d obtain i where i: "i \<in> ?U" "row i A \<in> span (rows A - {row i A})"
+  let ?U = "UNIV :: 'n set"
+  from d obtain i where i: "row i A \<in> span (rows A - {row i A})"
     unfolding dependent_def rows_def by blast
-  {fix j k assume j: "j \<in>?U" and k: "k \<in> ?U" and jk: "j \<noteq> k"
+  {fix j k assume jk: "j \<noteq> k"
     and c: "row j A = row k A"
-    from det_identical_rows[OF j k jk c] have ?thesis .}
+    from det_identical_rows[OF jk c] have ?thesis .}
   moreover
-  {assume H: "\<And> i j. i\<in> ?U \<Longrightarrow> j \<in> ?U \<Longrightarrow> i \<noteq> j \<Longrightarrow> row i A \<noteq> row j A"
-    have th0: "- row i A \<in> span {row j A|j. j \<in> ?U \<and> j \<noteq> i}"
+  {assume H: "\<And> i j. i \<noteq> j \<Longrightarrow> row i A \<noteq> row j A"
+    have th0: "- row i A \<in> span {row j A|j. j \<noteq> i}"
       apply (rule span_neg)
       apply (rule set_rev_mp)
-      apply (rule i(2))
+      apply (rule i)
       apply (rule span_mono)
       using H i by (auto simp add: rows_def)
-    from det_row_span[OF i(1) th0]
+    from det_row_span[OF th0]
     have "det A = det (\<chi> k. if k = i then 0 *s 1 else row k A)"
       unfolding right_minus vector_smult_lzero ..
-    with det_row_mul[OF i(1), of "0::'a" "\<lambda>i. 1"]
+    with det_row_mul[of i "0::'a" "\<lambda>i. 1"]
     have "det A = 0" by simp}
   ultimately show ?thesis by blast
 qed
 
-lemma det_dependent_columns: assumes d: "dependent(columns (A::'a::ordered_idom^'n^'n))" shows "det A = 0"
+lemma det_dependent_columns: assumes d: "dependent(columns (A::'a::ordered_idom^'n^'n::finite))" shows "det A = 0"
 by (metis d det_dependent_rows rows_transp det_transp)
 
 (* ------------------------------------------------------------------------- *)
 (* Multilinearity and the multiplication formula.                            *)
 (* ------------------------------------------------------------------------- *)
 
-lemma Cart_lambda_cong: "(\<And>x. x \<in> {1 .. dimindex (UNIV :: 'n set)} \<Longrightarrow> f x = g x) \<Longrightarrow> (Cart_lambda f::'a^'n) = (Cart_lambda g :: 'a^'n)"
+lemma Cart_lambda_cong: "(\<And>x. f x = g x) \<Longrightarrow> (Cart_lambda f::'a^'n) = (Cart_lambda g :: 'a^'n)"
   apply (rule iffD1[OF Cart_lambda_unique]) by vector
 
 lemma det_linear_row_setsum:
-  assumes fS: "finite S" and k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}"
-  shows "det ((\<chi> i. if i = k then setsum (a i) S else c i)::'a::comm_ring_1^'n^'n) = setsum (\<lambda>j. det ((\<chi> i. if i = k then a  i j else c i)::'a^'n^'n)) S"
-  using k
+  assumes fS: "finite S"
+  shows "det ((\<chi> i. if i = k then setsum (a i) S else c i)::'a::comm_ring_1^'n^'n::finite) = setsum (\<lambda>j. det ((\<chi> i. if i = k then a  i j else c i)::'a^'n^'n)) S"
 proof(induct rule: finite_induct[OF fS])
-  case 1 thus ?case apply simp  unfolding setsum_empty det_row_0[OF k] ..
+  case 1 thus ?case apply simp  unfolding setsum_empty det_row_0[of k] ..
 next
   case (2 x F)
   then  show ?case by (simp add: det_row_add cong del: if_weak_cong)
@@ -588,91 +536,89 @@
 lemma eq_id_iff[simp]: "(\<forall>x. f x = x) = (f = id)" by (auto intro: ext)
 
 lemma det_linear_rows_setsum_lemma:
-  assumes fS: "finite S" and k: "k \<le> dimindex (UNIV :: 'n set)"
-  shows "det((\<chi> i. if i <= k then setsum (a i) S else c i):: 'a::comm_ring_1^'n^'n) =
-             setsum (\<lambda>f. det((\<chi> i. if i <= k then a i (f i) else c i)::'a^'n^'n))
-                 {f. (\<forall>i \<in> {1 .. k}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1..k} \<longrightarrow> f i = i)}"
-using k
-proof(induct k arbitrary: a c)
-  case 0
-  have th0: "\<And>x y. (\<chi> i. if i <= 0 then x i else y i) = (\<chi> i. y i)" by vector
-  from "0.prems"  show ?case unfolding th0 by simp
+  assumes fS: "finite S" and fT: "finite T"
+  shows "det((\<chi> i. if i \<in> T then setsum (a i) S else c i):: 'a::comm_ring_1^'n^'n::finite) =
+             setsum (\<lambda>f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n))
+                 {f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}"
+using fT
+proof(induct T arbitrary: a c set: finite)
+  case empty
+  have th0: "\<And>x y. (\<chi> i. if i \<in> {} then x i else y i) = (\<chi> i. y i)" by vector
+  from "empty.prems"  show ?case unfolding th0 by simp
 next
-  case (Suc k a c)
-  let ?F = "\<lambda>k. {f. (\<forall>i \<in> {1 .. k}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1..k} \<longrightarrow> f i = i)}"
-  let ?h = "\<lambda>(y::nat,g) i. if i = Suc k then y else g i"
-  let ?k = "\<lambda>h. (h(Suc k),(\<lambda>i. if i = Suc k then i else h i))"
-  let ?s = "\<lambda> k a c f. det((\<chi> i. if i <= k then a i (f i) else c i)::'a^'n^'n)"
-  let ?c = "\<lambda>i. if i = Suc k then a i j else c i"
-  from Suc.prems have Sk: "Suc k \<in> {1 .. dimindex (UNIV :: 'n set)}" by simp
-  from Suc.prems have k': "k \<le> dimindex (UNIV :: 'n set)" by arith
-  have thif: "\<And>a b c d. (if b \<or> a then c else d) = (if a then c else if b then c else d)" by simp
+  case (insert z T a c)
+  let ?F = "\<lambda>T. {f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}"
+  let ?h = "\<lambda>(y,g) i. if i = z then y else g i"
+  let ?k = "\<lambda>h. (h(z),(\<lambda>i. if i = z then i else h i))"
+  let ?s = "\<lambda> k a c f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n)"
+  let ?c = "\<lambda>i. if i = z then a i j else c i"
+  have thif: "\<And>a b c d. (if a \<or> b then c else d) = (if a then c else if b then c else d)" by simp
   have thif2: "\<And>a b c d e. (if a then b else if c then d else e) =
      (if c then (if a then b else d) else (if a then b else e))" by simp
-  have "det (\<chi> i. if i \<le> Suc k then setsum (a i) S else c i) =
-        det (\<chi> i. if i = Suc k then setsum (a i) S
-                 else if i \<le> k then setsum (a i) S else c i)"
-    unfolding le_Suc_eq thif  ..
-  also have "\<dots> = (\<Sum>j\<in>S. det (\<chi> i. if i \<le> k then setsum (a i) S
-                    else if i = Suc k then a i j else c i))"
-    unfolding det_linear_row_setsum[OF fS Sk]
+  from `z \<notin> T` have nz: "\<And>i. i \<in> T \<Longrightarrow> i = z \<longleftrightarrow> False" by auto
+  have "det (\<chi> i. if i \<in> insert z T then setsum (a i) S else c i) =
+        det (\<chi> i. if i = z then setsum (a i) S
+                 else if i \<in> T then setsum (a i) S else c i)"
+    unfolding insert_iff thif ..
+  also have "\<dots> = (\<Sum>j\<in>S. det (\<chi> i. if i \<in> T then setsum (a i) S
+                    else if i = z then a i j else c i))"
+    unfolding det_linear_row_setsum[OF fS]
     apply (subst thif2)
-    by (simp cong del: if_weak_cong cong add: if_cong)
+    using nz by (simp cong del: if_weak_cong cong add: if_cong)
   finally have tha:
-    "det (\<chi> i. if i \<le> Suc k then setsum (a i) S else c i) =
-     (\<Sum>(j, f)\<in>S \<times> ?F k. det (\<chi> i. if i \<le> k then a i (f i)
-                                else if i = Suc k then a i j
+    "det (\<chi> i. if i \<in> insert z T then setsum (a i) S else c i) =
+     (\<Sum>(j, f)\<in>S \<times> ?F T. det (\<chi> i. if i \<in> T then a i (f i)
+                                else if i = z then a i j
                                 else c i))"
-    unfolding  Suc.hyps[OF k'] unfolding setsum_cartesian_product by blast
+    unfolding  insert.hyps unfolding setsum_cartesian_product by blast
   show ?case unfolding tha
     apply(rule setsum_eq_general_reverses[where h= "?h" and k= "?k"],
-      blast intro: finite_cartesian_product fS finite_bounded_functions[OF fS],
-      blast intro: finite_cartesian_product fS finite_bounded_functions[OF fS], auto intro: ext)
+      blast intro: finite_cartesian_product fS finite,
+      blast intro: finite_cartesian_product fS finite)
+    using `z \<notin> T`
+    apply (auto intro: ext)
     apply (rule cong[OF refl[of det]])
     by vector
 qed
 
 lemma det_linear_rows_setsum:
-  assumes fS: "finite S"
-  shows "det (\<chi> i. setsum (a i) S) = setsum (\<lambda>f. det (\<chi> i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n)) {f. (\<forall>i \<in> {1 .. dimindex (UNIV :: 'n set)}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1.. dimindex (UNIV :: 'n set)} \<longrightarrow> f i = i)}"
+  assumes fS: "finite (S::'n::finite set)"
+  shows "det (\<chi> i. setsum (a i) S) = setsum (\<lambda>f. det (\<chi> i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n::finite)) {f. \<forall>i. f i \<in> S}"
 proof-
-  have th0: "\<And>x y. ((\<chi> i. if i <= dimindex(UNIV:: 'n set) then x i else y i) :: 'a^'n^'n) = (\<chi> i. x i)" by vector
+  have th0: "\<And>x y. ((\<chi> i. if i \<in> (UNIV:: 'n set) then x i else y i) :: 'a^'n^'n) = (\<chi> i. x i)" by vector
 
-  from det_linear_rows_setsum_lemma[OF fS, of "dimindex (UNIV :: 'n set)" a, unfolded th0, OF order_refl] show ?thesis by blast
+  from det_linear_rows_setsum_lemma[OF fS, of "UNIV :: 'n set" a, unfolded th0, OF finite] show ?thesis by simp
 qed
 
 lemma matrix_mul_setsum_alt:
-  fixes A B :: "'a::comm_ring_1^'n^'n"
-  shows "A ** B = (\<chi> i. setsum (\<lambda>k. A$i$k *s B $ k) {1 .. dimindex (UNIV :: 'n set)})"
+  fixes A B :: "'a::comm_ring_1^'n^'n::finite"
+  shows "A ** B = (\<chi> i. setsum (\<lambda>k. A$i$k *s B $ k) (UNIV :: 'n set))"
   by (vector matrix_matrix_mult_def setsum_component)
 
 lemma det_rows_mul:
-  "det((\<chi> i. c i *s a i)::'a::comm_ring_1^'n^'n) =
-  setprod (\<lambda>i. c i) {1..dimindex(UNIV:: 'n set)} * det((\<chi> i. a i)::'a^'n^'n)"
-proof (simp add: det_def Cart_lambda_beta' setsum_right_distrib vector_component cong add: setprod_cong del: One_nat_def, rule setsum_cong2)
-  let ?U = "{1 .. dimindex(UNIV :: 'n set)}"
+  "det((\<chi> i. c i *s a i)::'a::comm_ring_1^'n^'n::finite) =
+  setprod (\<lambda>i. c i) (UNIV:: 'n set) * det((\<chi> i. a i)::'a^'n^'n)"
+proof (simp add: det_def setsum_right_distrib cong add: setprod_cong, rule setsum_cong2)
+  let ?U = "UNIV :: 'n set"
   let ?PU = "{p. p permutes ?U}"
   fix p assume pU: "p \<in> ?PU"
   let ?s = "of_int (sign p)"
   from pU have p: "p permutes ?U" by blast
-  have "setprod (\<lambda>i. (c i *s a i) $ p i) ?U = setprod (\<lambda>i. c i * a i $ p i) ?U"
-    apply (rule setprod_cong, blast)
-    by (auto simp only: permutes_in_image[OF p] intro: vector_smult_component)
-  also have "\<dots> = setprod c ?U * setprod (\<lambda>i. a i $ p i) ?U"
+  have "setprod (\<lambda>i. c i * a i $ p i) ?U = setprod c ?U * setprod (\<lambda>i. a i $ p i) ?U"
     unfolding setprod_timesf ..
-  finally show "?s * (\<Prod>xa\<in>?U. (c xa *s a xa) $ p xa) =
+  then show "?s * (\<Prod>xa\<in>?U. c xa * a xa $ p xa) =
         setprod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))" by (simp add: ring_simps)
 qed
 
 lemma det_mul:
-  fixes A B :: "'a::ordered_idom^'n^'n"
+  fixes A B :: "'a::ordered_idom^'n^'n::finite"
   shows "det (A ** B) = det A * det B"
 proof-
-  let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
+  let ?U = "UNIV :: 'n set"
   let ?F = "{f. (\<forall>i\<in> ?U. f i \<in> ?U) \<and> (\<forall>i. i \<notin> ?U \<longrightarrow> f i = i)}"
   let ?PU = "{p. p permutes ?U}"
   have fU: "finite ?U" by simp
-  have fF: "finite ?F"  using finite_bounded_functions[OF fU] .
+  have fF: "finite ?F" by (rule finite)
   {fix p assume p: "p permutes ?U"
 
     have "p \<in> ?F" unfolding mem_Collect_eq permutes_in_image[OF p]
@@ -687,23 +633,21 @@
     let ?A = "(\<chi> i. A$i$f i *s B$f i) :: 'a^'n^'n"
     let ?B = "(\<chi> i. B$f i) :: 'a^'n^'n"
     {assume fni: "\<not> inj_on f ?U"
-      then obtain i j where ij: "i \<in> ?U" "j \<in> ?U" "f i = f j" "i \<noteq> j"
+      then obtain i j where ij: "f i = f j" "i \<noteq> j"
 	unfolding inj_on_def by blast
       from ij
       have rth: "row i ?B = row j ?B" by (vector row_def)
-      from det_identical_rows[OF ij(1,2,4) rth]
+      from det_identical_rows[OF ij(2) rth]
       have "det (\<chi> i. A$i$f i *s B$f i) = 0"
 	unfolding det_rows_mul by simp}
     moreover
     {assume fi: "inj_on f ?U"
       from f fi have fith: "\<And>i j. f i = f j \<Longrightarrow> i = j"
-	unfolding inj_on_def
-	apply (case_tac "i \<in> ?U")
-	apply (case_tac "j \<in> ?U") by metis+
+	unfolding inj_on_def by metis
       note fs = fi[unfolded surjective_iff_injective_gen[OF fU fU refl fUU, symmetric]]
 
       {fix y
-	from fs f have "\<exists>x. f x = y" by (cases "y \<in> ?U") blast+
+	from fs f have "\<exists>x. f x = y" by blast
 	then obtain x where x: "f x = y" by blast
 	{fix z assume z: "f z = y" from fith x z have "z = x" by metis}
 	with x have "\<exists>!x. f x = y" by blast}
@@ -747,7 +691,7 @@
     unfolding det_def setsum_product
     by (rule setsum_cong2)
   have "det (A**B) = setsum (\<lambda>f.  det (\<chi> i. A $ i $ f i *s B $ f i)) ?F"
-    unfolding matrix_mul_setsum_alt det_linear_rows_setsum[OF fU] ..
+    unfolding matrix_mul_setsum_alt det_linear_rows_setsum[OF fU] by simp
   also have "\<dots> = setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU"
     using setsum_mono_zero_cong_left[OF fF PUF zth, symmetric]
     unfolding det_rows_mul by auto
@@ -759,17 +703,17 @@
 (* ------------------------------------------------------------------------- *)
 
 lemma invertible_left_inverse:
-  fixes A :: "real^'n^'n"
+  fixes A :: "real^'n^'n::finite"
   shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). B** A = mat 1)"
   by (metis invertible_def matrix_left_right_inverse)
 
 lemma invertible_righ_inverse:
-  fixes A :: "real^'n^'n"
+  fixes A :: "real^'n^'n::finite"
   shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). A** B = mat 1)"
   by (metis invertible_def matrix_left_right_inverse)
 
 lemma invertible_det_nz:
-  fixes A::"real ^'n^'n"
+  fixes A::"real ^'n^'n::finite"
   shows "invertible A \<longleftrightarrow> det A \<noteq> 0"
 proof-
   {assume "invertible A"
@@ -780,7 +724,7 @@
       apply (simp add: det_mul det_I) by algebra }
   moreover
   {assume H: "\<not> invertible A"
-    let ?U = "{1 .. dimindex(UNIV :: 'n set)}"
+    let ?U = "UNIV :: 'n set"
     have fU: "finite ?U" by simp
     from H obtain c i where c: "setsum (\<lambda>i. c i *s row i A) ?U = 0"
       and iU: "i \<in> ?U" and ci: "c i \<noteq> 0"
@@ -794,11 +738,11 @@
     from c ci
     have thr0: "- row i A = setsum (\<lambda>j. (1/ c i) *s c j *s row j A) (?U - {i})"
       unfolding setsum_diff1'[OF fU iU] setsum_cmul
-      apply (simp add: field_simps)
+      apply -
       apply (rule vector_mul_lcancel_imp[OF ci])
       apply (auto simp add: vector_smult_assoc vector_smult_rneg field_simps)
       unfolding stupid ..
-    have thr: "- row i A \<in> span {row j A| j. j\<in> ?U \<and> j \<noteq> i}"
+    have thr: "- row i A \<in> span {row j A| j. j \<noteq> i}"
       unfolding thr0
       apply (rule span_setsum)
       apply simp
@@ -810,8 +754,8 @@
     let ?B = "(\<chi> k. if k = i then 0 else row k A) :: real ^'n^'n"
     have thrb: "row i ?B = 0" using iU by (vector row_def)
     have "det A = 0"
-      unfolding det_row_span[OF iU thr, symmetric] right_minus
-      unfolding  det_zero_row[OF iU thrb]  ..}
+      unfolding det_row_span[OF thr, symmetric] right_minus
+      unfolding  det_zero_row[OF thrb]  ..}
   ultimately show ?thesis by blast
 qed
 
@@ -820,15 +764,14 @@
 (* ------------------------------------------------------------------------- *)
 
 lemma cramer_lemma_transp:
-  fixes A:: "'a::ordered_idom^'n^'n" and x :: "'a ^'n"
-  assumes k: "k \<in> {1 .. dimindex(UNIV ::'n set)}"
-  shows "det ((\<chi> i. if i = k then setsum (\<lambda>i. x$i *s row i A) {1 .. dimindex(UNIV::'n set)}
+  fixes A:: "'a::ordered_idom^'n^'n::finite" and x :: "'a ^'n::finite"
+  shows "det ((\<chi> i. if i = k then setsum (\<lambda>i. x$i *s row i A) (UNIV::'n set)
                            else row i A)::'a^'n^'n) = x$k * det A"
   (is "?lhs = ?rhs")
 proof-
-  let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
+  let ?U = "UNIV :: 'n set"
   let ?Uk = "?U - {k}"
-  have U: "?U = insert k ?Uk" using k by blast
+  have U: "?U = insert k ?Uk" by blast
   have fUk: "finite ?Uk" by simp
   have kUk: "k \<notin> ?Uk" by simp
   have th00: "\<And>k s. x$k *s row k A + s = (x$k - 1) *s row k A + row k A + s"
@@ -837,7 +780,7 @@
   have "(\<chi> i. row i A) = A" by (vector row_def)
   then have thd1: "det (\<chi> i. row i A) = det A"  by simp
   have thd0: "det (\<chi> i. if i = k then row k A + (\<Sum>i \<in> ?Uk. x $ i *s row i A) else row i A) = det A"
-    apply (rule det_row_span[OF k])
+    apply (rule det_row_span)
     apply (rule span_setsum[OF fUk])
     apply (rule ballI)
     apply (rule span_mul)
@@ -849,30 +792,29 @@
     unfolding setsum_insert[OF fUk kUk]
     apply (subst th00)
     unfolding add_assoc
-    apply (subst det_row_add[OF k])
+    apply (subst det_row_add)
     unfolding thd0
-    unfolding det_row_mul[OF k]
+    unfolding det_row_mul
     unfolding th001[of k "\<lambda>i. row i A"]
     unfolding thd1  by (simp add: ring_simps)
 qed
 
 lemma cramer_lemma:
-  fixes A :: "'a::ordered_idom ^'n^'n"
-  assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}" (is " _ \<in> ?U")
+  fixes A :: "'a::ordered_idom ^'n^'n::finite"
   shows "det((\<chi> i j. if j = k then (A *v x)$i else A$i$j):: 'a^'n^'n) = x$k * det A"
 proof-
+  let ?U = "UNIV :: 'n set"
   have stupid: "\<And>c. setsum (\<lambda>i. c i *s row i (transp A)) ?U = setsum (\<lambda>i. c i *s column i A) ?U"
     by (auto simp add: row_transp intro: setsum_cong2)
-  show ?thesis
-  unfolding matrix_mult_vsum
-  unfolding cramer_lemma_transp[OF k, of x "transp A", unfolded det_transp, symmetric]
+  show ?thesis  unfolding matrix_mult_vsum
+  unfolding cramer_lemma_transp[of k x "transp A", unfolded det_transp, symmetric]
   unfolding stupid[of "\<lambda>i. x$i"]
   apply (subst det_transp[symmetric])
   apply (rule cong[OF refl[of det]]) by (vector transp_def column_def row_def)
 qed
 
 lemma cramer:
-  fixes A ::"real^'n^'n"
+  fixes A ::"real^'n^'n::finite"
   assumes d0: "det A \<noteq> 0"
   shows "A *v x = b \<longleftrightarrow> x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)"
 proof-
@@ -884,7 +826,7 @@
   {fix x assume x: "A *v x = b"
   have "x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)"
     unfolding x[symmetric]
-    using d0 by (simp add: Cart_eq Cart_lambda_beta' cramer_lemma field_simps)}
+    using d0 by (simp add: Cart_eq cramer_lemma field_simps)}
   with xe show ?thesis by auto
 qed
 
@@ -894,7 +836,7 @@
 
 definition "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v w. f v \<bullet> f w = v \<bullet> w)"
 
-lemma orthogonal_transformation: "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>(v::real ^'n). norm (f v) = norm v)"
+lemma orthogonal_transformation: "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>(v::real ^_). norm (f v) = norm v)"
   unfolding orthogonal_transformation_def
   apply auto
   apply (erule_tac x=v in allE)+
@@ -903,14 +845,14 @@
 
 definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow> transp Q ** Q = mat 1 \<and> Q ** transp Q = mat 1"
 
-lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n)  \<longleftrightarrow> transp Q ** Q = mat 1"
+lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n::finite)  \<longleftrightarrow> transp Q ** Q = mat 1"
   by (metis matrix_left_right_inverse orthogonal_matrix_def)
 
-lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1)"
+lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n::finite)"
   by (simp add: orthogonal_matrix_def transp_mat matrix_mul_lid)
 
 lemma orthogonal_matrix_mul:
-  fixes A :: "real ^'n^'n"
+  fixes A :: "real ^'n^'n::finite"
   assumes oA : "orthogonal_matrix A"
   and oB: "orthogonal_matrix B"
   shows "orthogonal_matrix(A ** B)"
@@ -921,26 +863,26 @@
   by (simp add: matrix_mul_rid)
 
 lemma orthogonal_transformation_matrix:
-  fixes f:: "real^'n \<Rightarrow> real^'n"
+  fixes f:: "real^'n \<Rightarrow> real^'n::finite"
   shows "orthogonal_transformation f \<longleftrightarrow> linear f \<and> orthogonal_matrix(matrix f)"
   (is "?lhs \<longleftrightarrow> ?rhs")
 proof-
   let ?mf = "matrix f"
   let ?ot = "orthogonal_transformation f"
-  let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
+  let ?U = "UNIV :: 'n set"
   have fU: "finite ?U" by simp
   let ?m1 = "mat 1 :: real ^'n^'n"
   {assume ot: ?ot
     from ot have lf: "linear f" and fd: "\<forall>v w. f v \<bullet> f w = v \<bullet> w"
       unfolding  orthogonal_transformation_def orthogonal_matrix by blast+
-    {fix i j assume i: "i \<in> ?U" and j: "j \<in> ?U"
+    {fix i j
       let ?A = "transp ?mf ** ?mf"
       have th0: "\<And>b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)"
 	"\<And>b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)"
 	by simp_all
-      from fd[rule_format, of "basis i" "basis j", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul] i j
+      from fd[rule_format, of "basis i" "basis j", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul]
       have "?A$i$j = ?m1 $ i $ j"
-	by (simp add: Cart_lambda_beta' dot_def matrix_matrix_mult_def columnvector_def rowvector_def basis_def th0 setsum_delta[OF fU] mat_def del: One_nat_def)}
+	by (simp add: dot_def matrix_matrix_mult_def columnvector_def rowvector_def basis_def th0 setsum_delta[OF fU] mat_def)}
     hence "orthogonal_matrix ?mf" unfolding orthogonal_matrix by vector
     with lf have ?rhs by blast}
   moreover
@@ -949,12 +891,12 @@
       unfolding orthogonal_matrix_def norm_eq orthogonal_transformation
       unfolding matrix_works[OF lf, symmetric]
       apply (subst dot_matrix_vector_mul)
-      by (simp add: dot_matrix_product matrix_mul_lid del: One_nat_def)}
+      by (simp add: dot_matrix_product matrix_mul_lid)}
   ultimately show ?thesis by blast
 qed
 
 lemma det_orthogonal_matrix:
-  fixes Q:: "'a::ordered_idom^'n^'n"
+  fixes Q:: "'a::ordered_idom^'n^'n::finite"
   assumes oQ: "orthogonal_matrix Q"
   shows "det Q = 1 \<or> det Q = - 1"
 proof-
@@ -979,7 +921,7 @@
 (* Linearity of scaling, and hence isometry, that preserves origin.          *)
 (* ------------------------------------------------------------------------- *)
 lemma scaling_linear:
-  fixes f :: "real ^'n \<Rightarrow> real ^'n"
+  fixes f :: "real ^'n \<Rightarrow> real ^'n::finite"
   assumes f0: "f 0 = 0" and fd: "\<forall>x y. dist (f x) (f y) = c * dist x y"
   shows "linear f"
 proof-
@@ -995,7 +937,7 @@
 qed
 
 lemma isometry_linear:
-  "f (0:: real^'n) = (0:: real^'n) \<Longrightarrow> \<forall>x y. dist(f x) (f y) = dist x y
+  "f (0:: real^'n) = (0:: real^'n::finite) \<Longrightarrow> \<forall>x y. dist(f x) (f y) = dist x y
         \<Longrightarrow> linear f"
 by (rule scaling_linear[where c=1]) simp_all
 
@@ -1004,7 +946,7 @@
 (* ------------------------------------------------------------------------- *)
 
 lemma orthogonal_transformation_isometry:
-  "orthogonal_transformation f \<longleftrightarrow> f(0::real^'n) = (0::real^'n) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)"
+  "orthogonal_transformation f \<longleftrightarrow> f(0::real^'n) = (0::real^'n::finite) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)"
   unfolding orthogonal_transformation
   apply (rule iffI)
   apply clarify
@@ -1023,7 +965,7 @@
 (* ------------------------------------------------------------------------- *)
 
 lemma isometry_sphere_extend:
-  fixes f:: "real ^'n \<Rightarrow> real ^'n"
+  fixes f:: "real ^'n \<Rightarrow> real ^'n::finite"
   assumes f1: "\<forall>x. norm x = 1 \<longrightarrow> norm (f x) = 1"
   and fd1: "\<forall> x y. norm x = 1 \<longrightarrow> norm y = 1 \<longrightarrow> dist (f x) (f y) = dist x y"
   shows "\<exists>g. orthogonal_transformation g \<and> (\<forall>x. norm x = 1 \<longrightarrow> g x = f x)"
@@ -1095,7 +1037,7 @@
 definition "rotoinversion_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = - 1"
 
 lemma orthogonal_rotation_or_rotoinversion:
-  fixes Q :: "'a::ordered_idom^'n^'n"
+  fixes Q :: "'a::ordered_idom^'n^'n::finite"
   shows " orthogonal_matrix Q \<longleftrightarrow> rotation_matrix Q \<or> rotoinversion_matrix Q"
   by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix)
 (* ------------------------------------------------------------------------- *)
@@ -1110,17 +1052,16 @@
   by (simp add: nat_number setprod_numseg mult_commute)
 
 lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1"
-  by (simp add: det_def dimindex_def permutes_sing sign_id del: One_nat_def)
+  by (simp add: det_def permutes_sing sign_id UNIV_1)
 
 lemma det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1"
 proof-
-  have f12: "finite {2::nat}" "1 \<notin> {2::nat}" by auto
-  have th12: "{1 .. 2} = insert (1::nat) {2}" by auto
+  have f12: "finite {2::2}" "1 \<notin> {2::2}" by auto
   show ?thesis
-  apply (simp add: det_def dimindex_def th12 del: One_nat_def)
+  unfolding det_def UNIV_2
   unfolding setsum_over_permutations_insert[OF f12]
   unfolding permutes_sing
-  apply (simp add: sign_swap_id sign_id swap_id_eq del: One_nat_def)
+  apply (simp add: sign_swap_id sign_id swap_id_eq)
   by (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31))
 qed
 
@@ -1132,18 +1073,17 @@
   A$1$2 * A$2$1 * A$3$3 -
   A$1$3 * A$2$2 * A$3$1"
 proof-
-  have f123: "finite {(2::nat), 3}" "1 \<notin> {(2::nat), 3}" by auto
-  have f23: "finite {(3::nat)}" "2 \<notin> {(3::nat)}" by auto
-  have th12: "{1 .. 3} = insert (1::nat) (insert 2 {3})" by auto
+  have f123: "finite {2::3, 3}" "1 \<notin> {2::3, 3}" by auto
+  have f23: "finite {3::3}" "2 \<notin> {3::3}" by auto
 
   show ?thesis
-  apply (simp add: det_def dimindex_def th12 del: One_nat_def)
+  unfolding det_def UNIV_3
   unfolding setsum_over_permutations_insert[OF f123]
   unfolding setsum_over_permutations_insert[OF f23]
 
   unfolding permutes_sing
-  apply (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq del: One_nat_def)
-  apply (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31) One_nat_def)
+  apply (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq)
+  apply (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31))
   by (simp add: ring_simps)
 qed
 
--- a/src/HOL/Library/Efficient_Nat.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Efficient_Nat.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Implementation of natural numbers by target-language integers *}
 
 theory Efficient_Nat
-imports Code_Index Code_Integer
+imports Code_Index Code_Integer Main
 begin
 
 text {*
--- a/src/HOL/Library/Enum.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Enum.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Finite types as explicit enumerations *}
 
 theory Enum
-imports Plain "~~/src/HOL/Map"
+imports Map Main
 begin
 
 subsection {* Class @{text enum} *}
--- a/src/HOL/Library/Euclidean_Space.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Euclidean_Space.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,40 +5,59 @@
 header {* (Real) Vectors in Euclidean space, and elementary linear algebra.*}
 
 theory Euclidean_Space
-  imports "~~/src/HOL/Decision_Procs/Dense_Linear_Order" Complex_Main
+imports
+  Complex_Main "~~/src/HOL/Decision_Procs/Dense_Linear_Order"
   Finite_Cartesian_Product Glbs Infinite_Set Numeral_Type
   Inner_Product
-  uses ("normarith.ML")
+uses ("normarith.ML")
 begin
 
 text{* Some common special cases.*}
 
-lemma forall_1: "(\<forall>(i::'a::{order,one}). 1 <= i \<and> i <= 1 --> P i) \<longleftrightarrow> P 1"
-  by (metis order_eq_iff)
-lemma forall_dimindex_1: "(\<forall>i \<in> {1..dimindex(UNIV:: 1 set)}. P i) \<longleftrightarrow> P 1"
-  by (simp add: dimindex_def)
-
-lemma forall_2: "(\<forall>(i::nat). 1 <= i \<and> i <= 2 --> P i) \<longleftrightarrow> P 1 \<and> P 2"
-proof-
-  have "\<And>i::nat. 1 <= i \<and> i <= 2 \<longleftrightarrow> i = 1 \<or> i = 2" by arith
-  thus ?thesis by metis
+lemma forall_1: "(\<forall>i::1. P i) \<longleftrightarrow> P 1"
+  by (metis num1_eq_iff)
+
+lemma exhaust_2:
+  fixes x :: 2 shows "x = 1 \<or> x = 2"
+proof (induct x)
+  case (of_int z)
+  then have "0 <= z" and "z < 2" by simp_all
+  then have "z = 0 | z = 1" by arith
+  then show ?case by auto
 qed
 
-lemma forall_3: "(\<forall>(i::nat). 1 <= i \<and> i <= 3 --> P i) \<longleftrightarrow> P 1 \<and> P 2 \<and> P 3"
-proof-
-  have "\<And>i::nat. 1 <= i \<and> i <= 3 \<longleftrightarrow> i = 1 \<or> i = 2 \<or> i = 3" by arith
-  thus ?thesis by metis
+lemma forall_2: "(\<forall>i::2. P i) \<longleftrightarrow> P 1 \<and> P 2"
+  by (metis exhaust_2)
+
+lemma exhaust_3:
+  fixes x :: 3 shows "x = 1 \<or> x = 2 \<or> x = 3"
+proof (induct x)
+  case (of_int z)
+  then have "0 <= z" and "z < 3" by simp_all
+  then have "z = 0 \<or> z = 1 \<or> z = 2" by arith
+  then show ?case by auto
 qed
 
-lemma setsum_singleton[simp]: "setsum f {x} = f x" by simp
-lemma setsum_1: "setsum f {(1::'a::{order,one})..1} = f 1"
-  by (simp add: atLeastAtMost_singleton)
-
-lemma setsum_2: "setsum f {1::nat..2} = f 1 + f 2"
-  by (simp add: nat_number  atLeastAtMostSuc_conv add_commute)
-
-lemma setsum_3: "setsum f {1::nat..3} = f 1 + f 2 + f 3"
-  by (simp add: nat_number  atLeastAtMostSuc_conv add_commute)
+lemma forall_3: "(\<forall>i::3. P i) \<longleftrightarrow> P 1 \<and> P 2 \<and> P 3"
+  by (metis exhaust_3)
+
+lemma UNIV_1: "UNIV = {1::1}"
+  by (auto simp add: num1_eq_iff)
+
+lemma UNIV_2: "UNIV = {1::2, 2::2}"
+  using exhaust_2 by auto
+
+lemma UNIV_3: "UNIV = {1::3, 2::3, 3::3}"
+  using exhaust_3 by auto
+
+lemma setsum_1: "setsum f (UNIV::1 set) = f 1"
+  unfolding UNIV_1 by simp
+
+lemma setsum_2: "setsum f (UNIV::2 set) = f 1 + f 2"
+  unfolding UNIV_2 by simp
+
+lemma setsum_3: "setsum f (UNIV::3 set) = f 1 + f 2 + f 3"
+  unfolding UNIV_3 by (simp add: add_ac)
 
 subsection{* Basic componentwise operations on vectors. *}
 
@@ -76,10 +95,8 @@
 instantiation "^" :: (ord,type) ord
  begin
 definition vector_less_eq_def:
-  "less_eq (x :: 'a ^'b) y = (ALL i : {1 .. dimindex (UNIV :: 'b set)}.
-  x$i <= y$i)"
-definition vector_less_def: "less (x :: 'a ^'b) y = (ALL i : {1 ..
-  dimindex (UNIV :: 'b set)}. x$i < y$i)"
+  "less_eq (x :: 'a ^'b) y = (ALL i. x$i <= y$i)"
+definition vector_less_def: "less (x :: 'a ^'b) y = (ALL i. x$i < y$i)"
 
 instance by (intro_classes)
 end
@@ -102,19 +119,19 @@
 text{* Dot products. *}
 
 definition dot :: "'a::{comm_monoid_add, times} ^ 'n \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a" (infix "\<bullet>" 70) where
-  "x \<bullet> y = setsum (\<lambda>i. x$i * y$i) {1 .. dimindex (UNIV:: 'n set)}"
+  "x \<bullet> y = setsum (\<lambda>i. x$i * y$i) UNIV"
+
 lemma dot_1[simp]: "(x::'a::{comm_monoid_add, times}^1) \<bullet> y = (x$1) * (y$1)"
-  by (simp add: dot_def dimindex_def)
+  by (simp add: dot_def setsum_1)
 
 lemma dot_2[simp]: "(x::'a::{comm_monoid_add, times}^2) \<bullet> y = (x$1) * (y$1) + (x$2) * (y$2)"
-  by (simp add: dot_def dimindex_def nat_number)
+  by (simp add: dot_def setsum_2)
 
 lemma dot_3[simp]: "(x::'a::{comm_monoid_add, times}^3) \<bullet> y = (x$1) * (y$1) + (x$2) * (y$2) + (x$3) * (y$3)"
-  by (simp add: dot_def dimindex_def nat_number)
+  by (simp add: dot_def setsum_3)
 
 subsection {* A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space. *}
 
-lemmas Cart_lambda_beta' = Cart_lambda_beta[rule_format]
 method_setup vector = {*
 let
   val ss1 = HOL_basic_ss addsimps [@{thm dot_def}, @{thm setsum_addf} RS sym,
@@ -125,7 +142,7 @@
               @{thm vector_minus_def}, @{thm vector_uminus_def},
               @{thm vector_one_def}, @{thm vector_zero_def}, @{thm vec_def},
               @{thm vector_scaleR_def},
-              @{thm Cart_lambda_beta'}, @{thm vector_scalar_mult_def}]
+              @{thm Cart_lambda_beta}, @{thm vector_scalar_mult_def}]
  fun vector_arith_tac ths =
    simp_tac ss1
    THEN' (fn i => rtac @{thm setsum_cong2} i
@@ -145,39 +162,38 @@
 
 text{* Obvious "component-pushing". *}
 
-lemma vec_component: " i \<in> {1 .. dimindex (UNIV :: 'n set)} \<Longrightarrow> (vec x :: 'a ^ 'n)$i = x"
+lemma vec_component [simp]: "(vec x :: 'a ^ 'n)$i = x"
   by (vector vec_def)
 
-lemma vector_add_component:
-  fixes x y :: "'a::{plus} ^ 'n"  assumes i: "i \<in> {1 .. dimindex(UNIV:: 'n set)}"
+lemma vector_add_component [simp]:
+  fixes x y :: "'a::{plus} ^ 'n"
   shows "(x + y)$i = x$i + y$i"
-  using i by vector
-
-lemma vector_minus_component:
-  fixes x y :: "'a::{minus} ^ 'n"  assumes i: "i \<in> {1 .. dimindex(UNIV:: 'n set)}"
+  by vector
+
+lemma vector_minus_component [simp]:
+  fixes x y :: "'a::{minus} ^ 'n"
   shows "(x - y)$i = x$i - y$i"
-  using i  by vector
-
-lemma vector_mult_component:
-  fixes x y :: "'a::{times} ^ 'n"  assumes i: "i \<in> {1 .. dimindex(UNIV:: 'n set)}"
+  by vector
+
+lemma vector_mult_component [simp]:
+  fixes x y :: "'a::{times} ^ 'n"
   shows "(x * y)$i = x$i * y$i"
-  using i by vector
-
-lemma vector_smult_component:
-  fixes y :: "'a::{times} ^ 'n"  assumes i: "i \<in> {1 .. dimindex(UNIV:: 'n set)}"
+  by vector
+
+lemma vector_smult_component [simp]:
+  fixes y :: "'a::{times} ^ 'n"
   shows "(c *s y)$i = c * (y$i)"
-  using i by vector
-
-lemma vector_uminus_component:
-  fixes x :: "'a::{uminus} ^ 'n"  assumes i: "i \<in> {1 .. dimindex(UNIV:: 'n set)}"
+  by vector
+
+lemma vector_uminus_component [simp]:
+  fixes x :: "'a::{uminus} ^ 'n"
   shows "(- x)$i = - (x$i)"
-  using i by vector
-
-lemma vector_scaleR_component:
+  by vector
+
+lemma vector_scaleR_component [simp]:
   fixes x :: "'a::scaleR ^ 'n"
-  assumes i: "i \<in> {1 .. dimindex(UNIV :: 'n set)}"
   shows "(scaleR r x)$i = scaleR r (x$i)"
-  using i by vector
+  by vector
 
 lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector
 
@@ -250,7 +266,7 @@
 instance "^" :: (semiring_0,type) semiring_0
   apply (intro_classes) by (vector ring_simps)+
 instance "^" :: (semiring_1,type) semiring_1
-  apply (intro_classes) apply vector using dimindex_ge_1 by auto
+  apply (intro_classes) by vector
 instance "^" :: (comm_semiring,type) comm_semiring
   apply (intro_classes) by (vector ring_simps)+
 
@@ -274,16 +290,16 @@
 instance "^" :: (real_algebra_1,type) real_algebra_1 ..
 
 lemma of_nat_index:
-  "i\<in>{1 .. dimindex (UNIV :: 'n set)} \<Longrightarrow> (of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n"
+  "(of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n"
   apply (induct n)
   apply vector
   apply vector
   done
 lemma zero_index[simp]:
-  "i\<in>{1 .. dimindex (UNIV :: 'n set)} \<Longrightarrow> (0 :: 'a::zero ^'n)$i = 0" by vector
+  "(0 :: 'a::zero ^'n)$i = 0" by vector
 
 lemma one_index[simp]:
-  "i\<in>{1 .. dimindex (UNIV :: 'n set)} \<Longrightarrow> (1 :: 'a::one ^'n)$i = 1" by vector
+  "(1 :: 'a::one ^'n)$i = 1" by vector
 
 lemma one_plus_of_nat_neq_0: "(1::'a::semiring_char_0) + of_nat n \<noteq> 0"
 proof-
@@ -296,28 +312,7 @@
 proof (intro_classes)
   fix m n ::nat
   show "(of_nat m :: 'a^'b) = of_nat n \<longleftrightarrow> m = n"
-  proof(induct m arbitrary: n)
-    case 0 thus ?case apply vector
-      apply (induct n,auto simp add: ring_simps)
-      using dimindex_ge_1 apply auto
-      apply vector
-      by (auto simp add: of_nat_index one_plus_of_nat_neq_0)
-  next
-    case (Suc n m)
-    thus ?case  apply vector
-      apply (induct m, auto simp add: ring_simps of_nat_index zero_index)
-      using dimindex_ge_1 apply simp apply blast
-      apply (simp add: one_plus_of_nat_neq_0)
-      using dimindex_ge_1 apply simp apply blast
-      apply (simp add: vector_component one_index of_nat_index)
-      apply (simp only: of_nat.simps(2)[where ?'a = 'a, symmetric] of_nat_eq_iff)
-      using  dimindex_ge_1 apply simp apply blast
-      apply (simp add: vector_component one_index of_nat_index)
-      apply (simp only: of_nat.simps(2)[where ?'a = 'a, symmetric] of_nat_eq_iff)
-      using dimindex_ge_1 apply simp apply blast
-      apply (simp add: vector_component one_index of_nat_index)
-      done
-  qed
+    by (simp add: Cart_eq of_nat_index)
 qed
 
 instance "^" :: (comm_ring_1,type) comm_ring_1 by intro_classes
@@ -341,8 +336,7 @@
   by (vector ring_simps)
 
 lemma vec_eq[simp]: "(vec m = vec n) \<longleftrightarrow> (m = n)"
-  apply (auto simp add: vec_def Cart_eq vec_component Cart_lambda_beta )
-  using dimindex_ge_1 apply auto done
+  by (simp add: Cart_eq)
 
 subsection {* Square root of sum of squares *}
 
@@ -513,11 +507,11 @@
 
 subsection {* Norms *}
 
-instantiation "^" :: (real_normed_vector, type) real_normed_vector
+instantiation "^" :: (real_normed_vector, finite) real_normed_vector
 begin
 
 definition vector_norm_def:
-  "norm (x::'a^'b) = setL2 (\<lambda>i. norm (x$i)) {1 .. dimindex (UNIV:: 'b set)}"
+  "norm (x::'a^'b) = setL2 (\<lambda>i. norm (x$i)) UNIV"
 
 definition vector_sgn_def:
   "sgn (x::'a^'b) = scaleR (inverse (norm x)) x"
@@ -533,14 +527,11 @@
   show "norm (x + y) \<le> norm x + norm y"
     unfolding vector_norm_def
     apply (rule order_trans [OF _ setL2_triangle_ineq])
-    apply (rule setL2_mono)
-    apply (simp add: vector_component norm_triangle_ineq)
-    apply simp
+    apply (simp add: setL2_mono norm_triangle_ineq)
     done
   show "norm (scaleR a x) = \<bar>a\<bar> * norm x"
     unfolding vector_norm_def
-    by (simp add: vector_component norm_scaleR setL2_right_distrib
-             cong: strong_setL2_cong)
+    by (simp add: norm_scaleR setL2_right_distrib)
   show "sgn x = scaleR (inverse (norm x)) x"
     by (rule vector_sgn_def)
 qed
@@ -549,11 +540,11 @@
 
 subsection {* Inner products *}
 
-instantiation "^" :: (real_inner, type) real_inner
+instantiation "^" :: (real_inner, finite) real_inner
 begin
 
 definition vector_inner_def:
-  "inner x y = setsum (\<lambda>i. inner (x$i) (y$i)) {1 .. dimindex(UNIV::'b set)}"
+  "inner x y = setsum (\<lambda>i. inner (x$i) (y$i)) UNIV"
 
 instance proof
   fix r :: real and x y z :: "'a ^ 'b"
@@ -562,10 +553,10 @@
     by (simp add: inner_commute)
   show "inner (x + y) z = inner x z + inner y z"
     unfolding vector_inner_def
-    by (vector inner_left_distrib)
+    by (simp add: inner_left_distrib setsum_addf)
   show "inner (scaleR r x) y = r * inner x y"
     unfolding vector_inner_def
-    by (vector inner_scaleR_left)
+    by (simp add: inner_scaleR_left setsum_right_distrib)
   show "0 \<le> inner x x"
     unfolding vector_inner_def
     by (simp add: setsum_nonneg)
@@ -613,25 +604,16 @@
   show ?case by (simp add: h)
 qed
 
-lemma dot_eq_0: "x \<bullet> x = 0 \<longleftrightarrow> (x::'a::{ordered_ring_strict,ring_no_zero_divisors} ^ 'n) = 0"
-proof-
-  {assume f: "finite (UNIV :: 'n set)"
-    let ?S = "{Suc 0 .. card (UNIV :: 'n set)}"
-    have fS: "finite ?S" using f by simp
-    have fp: "\<forall> i\<in> ?S. x$i * x$i>= 0" by simp
-    have ?thesis by (vector dimindex_def f setsum_squares_eq_0_iff[OF fS fp])}
-  moreover
-  {assume "\<not> finite (UNIV :: 'n set)" then have ?thesis by (vector dimindex_def)}
-  ultimately show ?thesis by metis
-qed
-
-lemma dot_pos_lt[simp]: "(0 < x \<bullet> x) \<longleftrightarrow> (x::'a::{ordered_ring_strict,ring_no_zero_divisors} ^ 'n) \<noteq> 0" using dot_eq_0[of x] dot_pos_le[of x]
+lemma dot_eq_0: "x \<bullet> x = 0 \<longleftrightarrow> (x::'a::{ordered_ring_strict,ring_no_zero_divisors} ^ 'n::finite) = 0"
+  by (simp add: dot_def setsum_squares_eq_0_iff Cart_eq)
+
+lemma dot_pos_lt[simp]: "(0 < x \<bullet> x) \<longleftrightarrow> (x::'a::{ordered_ring_strict,ring_no_zero_divisors} ^ 'n::finite) \<noteq> 0" using dot_eq_0[of x] dot_pos_le[of x]
   by (auto simp add: le_less)
 
 subsection{* The collapse of the general concepts to dimension one. *}
 
 lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))"
-  by (vector dimindex_def)
+  by (simp add: Cart_eq forall_1)
 
 lemma forall_one: "(\<forall>(x::'a ^1). P x) \<longleftrightarrow> (\<forall>x. P(\<chi> i. x))"
   apply auto
@@ -640,7 +622,7 @@
   done
 
 lemma norm_vector_1: "norm (x :: _^1) = norm (x$1)"
-  by (simp add: vector_norm_def dimindex_def)
+  by (simp add: vector_norm_def UNIV_1)
 
 lemma norm_real: "norm(x::real ^ 1) = abs(x$1)"
   by (simp add: norm_vector_1)
@@ -648,17 +630,16 @@
 text{* Metric *}
 
 text {* FIXME: generalize to arbitrary @{text real_normed_vector} types *}
-definition dist:: "real ^ 'n \<Rightarrow> real ^ 'n \<Rightarrow> real" where
+definition dist:: "real ^ 'n::finite \<Rightarrow> real ^ 'n \<Rightarrow> real" where
   "dist x y = norm (x - y)"
 
 lemma dist_real: "dist(x::real ^ 1) y = abs((x$1) - (y$1))"
-  using dimindex_ge_1[of "UNIV :: 1 set"]
-  by (auto simp add: norm_real dist_def vector_component Cart_lambda_beta[where ?'a = "1"] )
+  by (auto simp add: norm_real dist_def)
 
 subsection {* A connectedness or intermediate value lemma with several applications. *}
 
 lemma connected_real_lemma:
-  fixes f :: "real \<Rightarrow> real ^ 'n"
+  fixes f :: "real \<Rightarrow> real ^ 'n::finite"
   assumes ab: "a \<le> b" and fa: "f a \<in> e1" and fb: "f b \<in> e2"
   and dst: "\<And>e x. a <= x \<Longrightarrow> x <= b \<Longrightarrow> 0 < e ==> \<exists>d > 0. \<forall>y. abs(y - x) < d \<longrightarrow> dist(f y) (f x) < e"
   and e1: "\<forall>y \<in> e1. \<exists>e > 0. \<forall>y'. dist y' y < e \<longrightarrow> y' \<in> e1"
@@ -758,7 +739,11 @@
 
 text{* Hence derive more interesting properties of the norm. *}
 
-lemma norm_0[simp]: "norm (0::real ^ 'n) = 0"
+text {*
+  This type-specific version is only here
+  to make @{text normarith.ML} happy.
+*}
+lemma norm_0: "norm (0::real ^ _) = 0"
   by (rule norm_zero)
 
 lemma norm_mul[simp]: "norm(a *s x) = abs(a) * norm x"
@@ -770,7 +755,7 @@
   by (simp add: vector_norm_def setL2_def dot_def power2_eq_square)
 lemma norm_pow_2: "norm x ^ 2 = x \<bullet> x"
   by (simp add: real_vector_norm_def)
-lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n)" by (metis norm_eq_zero)
+lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n::finite)" by (metis norm_eq_zero)
 lemma vector_mul_eq_0[simp]: "(a *s x = 0) \<longleftrightarrow> a = (0::'a::idom) \<or> x = 0"
   by vector
 lemma vector_mul_lcancel[simp]: "a *s x = a *s y \<longleftrightarrow> a = (0::real) \<or> x = y"
@@ -781,7 +766,9 @@
   by (metis vector_mul_lcancel)
 lemma vector_mul_rcancel_imp: "x \<noteq> 0 \<Longrightarrow> (a::real) *s x = b *s x ==> a = b"
   by (metis vector_mul_rcancel)
-lemma norm_cauchy_schwarz: "x \<bullet> y <= norm x * norm y"
+lemma norm_cauchy_schwarz:
+  fixes x y :: "real ^ 'n::finite"
+  shows "x \<bullet> y <= norm x * norm y"
 proof-
   {assume "norm x = 0"
     hence ?thesis by (simp add: dot_lzero dot_rzero)}
@@ -802,50 +789,74 @@
   ultimately show ?thesis by metis
 qed
 
-lemma norm_cauchy_schwarz_abs: "\<bar>x \<bullet> y\<bar> \<le> norm x * norm y"
+lemma norm_cauchy_schwarz_abs:
+  fixes x y :: "real ^ 'n::finite"
+  shows "\<bar>x \<bullet> y\<bar> \<le> norm x * norm y"
   using norm_cauchy_schwarz[of x y] norm_cauchy_schwarz[of x "-y"]
   by (simp add: real_abs_def dot_rneg)
 
-lemma norm_triangle_sub: "norm (x::real ^'n) <= norm(y) + norm(x - y)"
+lemma norm_triangle_sub: "norm (x::real ^'n::finite) <= norm(y) + norm(x - y)"
   using norm_triangle_ineq[of "y" "x - y"] by (simp add: ring_simps)
-lemma norm_triangle_le: "norm(x::real ^'n) + norm y <= e ==> norm(x + y) <= e"
+lemma norm_triangle_le: "norm(x::real ^'n::finite) + norm y <= e ==> norm(x + y) <= e"
   by (metis order_trans norm_triangle_ineq)
-lemma norm_triangle_lt: "norm(x::real ^'n) + norm(y) < e ==> norm(x + y) < e"
+lemma norm_triangle_lt: "norm(x::real ^'n::finite) + norm(y) < e ==> norm(x + y) < e"
   by (metis basic_trans_rules(21) norm_triangle_ineq)
 
-lemma component_le_norm: "i \<in> {1 .. dimindex(UNIV :: 'n set)} ==> \<bar>x$i\<bar> <= norm (x::real ^ 'n)"
+lemma setsum_delta:
+  assumes fS: "finite S"
+  shows "setsum (\<lambda>k. if k=a then b k else 0) S = (if a \<in> S then b a else 0)"
+proof-
+  let ?f = "(\<lambda>k. if k=a then b k else 0)"
+  {assume a: "a \<notin> S"
+    hence "\<forall> k\<in> S. ?f k = 0" by simp
+    hence ?thesis  using a by simp}
+  moreover
+  {assume a: "a \<in> S"
+    let ?A = "S - {a}"
+    let ?B = "{a}"
+    have eq: "S = ?A \<union> ?B" using a by blast
+    have dj: "?A \<inter> ?B = {}" by simp
+    from fS have fAB: "finite ?A" "finite ?B" by auto
+    have "setsum ?f S = setsum ?f ?A + setsum ?f ?B"
+      using setsum_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
+      by simp
+    then have ?thesis  using a by simp}
+  ultimately show ?thesis by blast
+qed
+
+lemma component_le_norm: "\<bar>x$i\<bar> <= norm (x::real ^ 'n::finite)"
   apply (simp add: vector_norm_def)
   apply (rule member_le_setL2, simp_all)
   done
 
-lemma norm_bound_component_le: "norm(x::real ^ 'n) <= e
-                ==> \<forall>i \<in> {1 .. dimindex(UNIV:: 'n set)}. \<bar>x$i\<bar> <= e"
+lemma norm_bound_component_le: "norm(x::real ^ 'n::finite) <= e
+                ==> \<bar>x$i\<bar> <= e"
   by (metis component_le_norm order_trans)
 
-lemma norm_bound_component_lt: "norm(x::real ^ 'n) < e
-                ==> \<forall>i \<in> {1 .. dimindex(UNIV:: 'n set)}. \<bar>x$i\<bar> < e"
+lemma norm_bound_component_lt: "norm(x::real ^ 'n::finite) < e
+                ==> \<bar>x$i\<bar> < e"
   by (metis component_le_norm basic_trans_rules(21))
 
-lemma norm_le_l1: "norm (x:: real ^'n) <= setsum(\<lambda>i. \<bar>x$i\<bar>) {1..dimindex(UNIV::'n set)}"
+lemma norm_le_l1: "norm (x:: real ^'n::finite) <= setsum(\<lambda>i. \<bar>x$i\<bar>) UNIV"
   by (simp add: vector_norm_def setL2_le_setsum)
 
-lemma real_abs_norm[simp]: "\<bar> norm x\<bar> = norm (x :: real ^'n)"
+lemma real_abs_norm: "\<bar>norm x\<bar> = norm (x :: real ^ _)"
   by (rule abs_norm_cancel)
-lemma real_abs_sub_norm: "\<bar>norm(x::real ^'n) - norm y\<bar> <= norm(x - y)"
+lemma real_abs_sub_norm: "\<bar>norm(x::real ^'n::finite) - norm y\<bar> <= norm(x - y)"
   by (rule norm_triangle_ineq3)
-lemma norm_le: "norm(x::real ^ 'n) <= norm(y) \<longleftrightarrow> x \<bullet> x <= y \<bullet> y"
+lemma norm_le: "norm(x::real ^ _) <= norm(y) \<longleftrightarrow> x \<bullet> x <= y \<bullet> y"
   by (simp add: real_vector_norm_def)
-lemma norm_lt: "norm(x::real ^'n) < norm(y) \<longleftrightarrow> x \<bullet> x < y \<bullet> y"
+lemma norm_lt: "norm(x::real ^ _) < norm(y) \<longleftrightarrow> x \<bullet> x < y \<bullet> y"
   by (simp add: real_vector_norm_def)
-lemma norm_eq: "norm (x::real ^'n) = norm y \<longleftrightarrow> x \<bullet> x = y \<bullet> y"
+lemma norm_eq: "norm (x::real ^ _) = norm y \<longleftrightarrow> x \<bullet> x = y \<bullet> y"
   by (simp add: order_eq_iff norm_le)
-lemma norm_eq_1: "norm(x::real ^ 'n) = 1 \<longleftrightarrow> x \<bullet> x = 1"
+lemma norm_eq_1: "norm(x::real ^ _) = 1 \<longleftrightarrow> x \<bullet> x = 1"
   by (simp add: real_vector_norm_def)
 
 text{* Squaring equations and inequalities involving norms.  *}
 
 lemma dot_square_norm: "x \<bullet> x = norm(x)^2"
-  by (simp add: real_vector_norm_def  dot_pos_le )
+  by (simp add: real_vector_norm_def)
 
 lemma norm_eq_square: "norm(x) = a \<longleftrightarrow> 0 <= a \<and> x \<bullet> x = a^2"
   by (auto simp add: real_vector_norm_def)
@@ -885,7 +896,7 @@
 
 text{* Equality of vectors in terms of @{term "op \<bullet>"} products.    *}
 
-lemma vector_eq: "(x:: real ^ 'n) = y \<longleftrightarrow> x \<bullet> x = x \<bullet> y\<and> y \<bullet> y = x \<bullet> x" (is "?lhs \<longleftrightarrow> ?rhs")
+lemma vector_eq: "(x:: real ^ 'n::finite) = y \<longleftrightarrow> x \<bullet> x = x \<bullet> y\<and> y \<bullet> y = x \<bullet> x" (is "?lhs \<longleftrightarrow> ?rhs")
 proof
   assume "?lhs" then show ?rhs by simp
 next
@@ -907,7 +918,7 @@
   done
 
   (* FIXME: Move all these theorems into the ML code using lemma antiquotation *)
-lemma norm_add_rule_thm: "b1 >= norm(x1 :: real ^'n) \<Longrightarrow> b2 >= norm(x2) ==> b1 + b2 >= norm(x1 + x2)"
+lemma norm_add_rule_thm: "b1 >= norm(x1 :: real ^'n::finite) \<Longrightarrow> b2 >= norm(x2) ==> b1 + b2 >= norm(x1 + x2)"
   apply (rule norm_triangle_le) by simp
 
 lemma ge_iff_diff_ge_0: "(a::'a::ordered_ring) \<ge> b == a - b \<ge> 0"
@@ -936,13 +947,13 @@
   "(c *s x + w) + (d *s y + z) == d *s y + ((c *s x + w) + z)" by ((atomize (full)), vector)+
 lemma pth_d: "x + (0::real ^'n) == x" by (atomize (full)) vector
 
-lemma norm_imp_pos_and_ge: "norm (x::real ^ 'n) == n \<Longrightarrow> norm x \<ge> 0 \<and> n \<ge> norm x"
+lemma norm_imp_pos_and_ge: "norm (x::real ^ _) == n \<Longrightarrow> norm x \<ge> 0 \<and> n \<ge> norm x"
   by (atomize) (auto simp add: norm_ge_zero)
 
 lemma real_eq_0_iff_le_ge_0: "(x::real) = 0 == x \<ge> 0 \<and> -x \<ge> 0" by arith
 
 lemma norm_pths:
-  "(x::real ^'n) = y \<longleftrightarrow> norm (x - y) \<le> 0"
+  "(x::real ^'n::finite) = y \<longleftrightarrow> norm (x - y) \<le> 0"
   "x \<noteq> y \<longleftrightarrow> \<not> (norm (x - y) \<le> 0)"
   using norm_ge_zero[of "x - y"] by auto
 
@@ -988,13 +999,13 @@
 
 lemma dist_le_0[simp]: "dist x y <= 0 \<longleftrightarrow> x = y" by norm
 
+lemma setsum_component [simp]:
+  fixes f:: " 'a \<Rightarrow> ('b::comm_monoid_add) ^'n"
+  shows "(setsum f S)$i = setsum (\<lambda>x. (f x)$i) S"
+  by (cases "finite S", induct S set: finite, simp_all)
+
 lemma setsum_eq: "setsum f S = (\<chi> i. setsum (\<lambda>x. (f x)$i ) S)"
-  apply vector
-  apply auto
-  apply (cases "finite S")
-  apply (rule finite_induct[of S])
-  apply (auto simp add: vector_component zero_index)
-  done
+  by (simp add: Cart_eq)
 
 lemma setsum_clauses:
   shows "setsum f {} = 0"
@@ -1005,13 +1016,7 @@
 lemma setsum_cmul:
   fixes f:: "'c \<Rightarrow> ('a::semiring_1)^'n"
   shows "setsum (\<lambda>x. c *s f x) S = c *s setsum f S"
-  by (simp add: setsum_eq Cart_eq Cart_lambda_beta vector_component setsum_right_distrib)
-
-lemma setsum_component:
-  fixes f:: " 'a \<Rightarrow> ('b::semiring_1) ^'n"
-  assumes i: "i \<in> {1 .. dimindex(UNIV:: 'n set)}"
-  shows "(setsum f S)$i = setsum (\<lambda>x. (f x)$i) S"
-  using i by (simp add: setsum_eq Cart_lambda_beta)
+  by (simp add: Cart_eq setsum_right_distrib)
 
 lemma setsum_norm:
   fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
@@ -1028,7 +1033,7 @@
 qed
 
 lemma real_setsum_norm:
-  fixes f :: "'a \<Rightarrow> real ^'n"
+  fixes f :: "'a \<Rightarrow> real ^'n::finite"
   assumes fS: "finite S"
   shows "norm (setsum f S) <= setsum (\<lambda>x. norm(f x)) S"
 proof(induct rule: finite_induct[OF fS])
@@ -1054,7 +1059,7 @@
 qed
 
 lemma real_setsum_norm_le:
-  fixes f :: "'a \<Rightarrow> real ^ 'n"
+  fixes f :: "'a \<Rightarrow> real ^ 'n::finite"
   assumes fS: "finite S"
   and fg: "\<forall>x \<in> S. norm (f x) \<le> g x"
   shows "norm (setsum f S) \<le> setsum g S"
@@ -1074,7 +1079,7 @@
   by simp
 
 lemma real_setsum_norm_bound:
-  fixes f :: "'a \<Rightarrow> real ^ 'n"
+  fixes f :: "'a \<Rightarrow> real ^ 'n::finite"
   assumes fS: "finite S"
   and K: "\<forall>x \<in> S. norm (f x) \<le> K"
   shows "norm (setsum f S) \<le> of_nat (card S) * K"
@@ -1155,13 +1160,13 @@
 by (auto intro: setsum_0')
 
 lemma vsum_norm_allsubsets_bound:
-  fixes f:: "'a \<Rightarrow> real ^'n"
+  fixes f:: "'a \<Rightarrow> real ^'n::finite"
   assumes fP: "finite P" and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (setsum f Q) \<le> e"
-  shows "setsum (\<lambda>x. norm (f x)) P \<le> 2 * real (dimindex(UNIV :: 'n set)) *  e"
+  shows "setsum (\<lambda>x. norm (f x)) P \<le> 2 * real CARD('n) *  e"
 proof-
-  let ?d = "real (dimindex (UNIV ::'n set))"
+  let ?d = "real CARD('n)"
   let ?nf = "\<lambda>x. norm (f x)"
-  let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
+  let ?U = "UNIV :: 'n set"
   have th0: "setsum (\<lambda>x. setsum (\<lambda>i. \<bar>f x $ i\<bar>) ?U) P = setsum (\<lambda>i. setsum (\<lambda>x. \<bar>f x $ i\<bar>) P) ?U"
     by (rule setsum_commute)
   have th1: "2 * ?d * e = of_nat (card ?U) * (2 * e)" by (simp add: real_of_nat_def)
@@ -1178,11 +1183,11 @@
     have thp0: "?Pp \<inter> ?Pn ={}" by auto
     have PpP: "?Pp \<subseteq> P" and PnP: "?Pn \<subseteq> P" by blast+
     have Ppe:"setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pp \<le> e"
-      using i component_le_norm[OF i, of "setsum (\<lambda>x. f x) ?Pp"]  fPs[OF PpP]
-      by (auto simp add: setsum_component intro: abs_le_D1)
+      using component_le_norm[of "setsum (\<lambda>x. f x) ?Pp" i]  fPs[OF PpP]
+      by (auto intro: abs_le_D1)
     have Pne: "setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pn \<le> e"
-      using i component_le_norm[OF i, of "setsum (\<lambda>x. - f x) ?Pn"]  fPs[OF PnP]
-      by (auto simp add: setsum_negf setsum_component vector_component intro: abs_le_D1)
+      using component_le_norm[of "setsum (\<lambda>x. - f x) ?Pn" i]  fPs[OF PnP]
+      by (auto simp add: setsum_negf intro: abs_le_D1)
     have "setsum (\<lambda>x. \<bar>f x $ i\<bar>) P = setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pp + setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pn"
       apply (subst thp)
       apply (rule setsum_Un_zero)
@@ -1204,32 +1209,29 @@
 
 definition "basis k = (\<chi> i. if i = k then 1 else 0)"
 
+lemma basis_component [simp]: "basis k $ i = (if k=i then 1 else 0)"
+  unfolding basis_def by simp
+
 lemma delta_mult_idempotent:
   "(if k=a then 1 else (0::'a::semiring_1)) * (if k=a then 1 else 0) = (if k=a then 1 else 0)" by (cases "k=a", auto)
 
 lemma norm_basis:
-  assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}"
-  shows "norm (basis k :: real ^'n) = 1"
-  using k
+  shows "norm (basis k :: real ^'n::finite) = 1"
   apply (simp add: basis_def real_vector_norm_def dot_def)
   apply (vector delta_mult_idempotent)
-  using setsum_delta[of "{1 .. dimindex (UNIV :: 'n set)}" "k" "\<lambda>k. 1::real"]
+  using setsum_delta[of "UNIV :: 'n set" "k" "\<lambda>k. 1::real"]
   apply auto
   done
 
-lemma norm_basis_1: "norm(basis 1 :: real ^'n) = 1"
-  apply (simp add: basis_def real_vector_norm_def dot_def)
-  apply (vector delta_mult_idempotent)
-  using setsum_delta[of "{1 .. dimindex (UNIV :: 'n set)}" "1" "\<lambda>k. 1::real"] dimindex_nonzero[of "UNIV :: 'n set"]
-  apply auto
-  done
-
-lemma vector_choose_size: "0 <= c ==> \<exists>(x::real^'n). norm x = c"
-  apply (rule exI[where x="c *s basis 1"])
-  by (simp only: norm_mul norm_basis_1)
+lemma norm_basis_1: "norm(basis 1 :: real ^'n::{finite,one}) = 1"
+  by (rule norm_basis)
+
+lemma vector_choose_size: "0 <= c ==> \<exists>(x::real^'n::finite). norm x = c"
+  apply (rule exI[where x="c *s basis arbitrary"])
+  by (simp only: norm_mul norm_basis)
 
 lemma vector_choose_dist: assumes e: "0 <= e"
-  shows "\<exists>(y::real^'n). dist x y = e"
+  shows "\<exists>(y::real^'n::finite). dist x y = e"
 proof-
   from vector_choose_size[OF e] obtain c:: "real ^'n"  where "norm c = e"
     by blast
@@ -1237,56 +1239,50 @@
   then show ?thesis by blast
 qed
 
-lemma basis_inj: "inj_on (basis :: nat \<Rightarrow> real ^'n) {1 .. dimindex (UNIV :: 'n set)}"
-  by (auto simp add: inj_on_def basis_def Cart_eq Cart_lambda_beta)
-
-lemma basis_component: "i \<in> {1 .. dimindex(UNIV:: 'n set)} ==> (basis k ::('a::semiring_1)^'n)$i = (if k=i then 1 else 0)"
-  by (simp add: basis_def Cart_lambda_beta)
+lemma basis_inj: "inj (basis :: 'n \<Rightarrow> real ^'n::finite)"
+  by (simp add: inj_on_def Cart_eq)
 
 lemma cond_value_iff: "f (if b then x else y) = (if b then f x else f y)"
   by auto
 
 lemma basis_expansion:
-  "setsum (\<lambda>i. (x$i) *s basis i) {1 .. dimindex (UNIV :: 'n set)} = (x::('a::ring_1) ^'n)" (is "?lhs = ?rhs" is "setsum ?f ?S = _")
-  by (auto simp add: Cart_eq basis_component[where ?'n = "'n"] setsum_component vector_component cond_value_iff setsum_delta[of "?S", where ?'b = "'a", simplified] cong del: if_weak_cong)
+  "setsum (\<lambda>i. (x$i) *s basis i) UNIV = (x::('a::ring_1) ^'n::finite)" (is "?lhs = ?rhs" is "setsum ?f ?S = _")
+  by (auto simp add: Cart_eq cond_value_iff setsum_delta[of "?S", where ?'b = "'a", simplified] cong del: if_weak_cong)
 
 lemma basis_expansion_unique:
-  "setsum (\<lambda>i. f i *s basis i) {1 .. dimindex (UNIV :: 'n set)} = (x::('a::comm_ring_1) ^'n) \<longleftrightarrow> (\<forall>i\<in>{1 .. dimindex(UNIV:: 'n set)}. f i = x$i)"
-  by (simp add: Cart_eq setsum_component vector_component basis_component setsum_delta cond_value_iff cong del: if_weak_cong)
+  "setsum (\<lambda>i. f i *s basis i) UNIV = (x::('a::comm_ring_1) ^'n::finite) \<longleftrightarrow> (\<forall>i. f i = x$i)"
+  by (simp add: Cart_eq setsum_delta cond_value_iff cong del: if_weak_cong)
 
 lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)"
   by auto
 
 lemma dot_basis:
-  assumes i: "i \<in> {1 .. dimindex (UNIV :: 'n set)}"
-  shows "basis i \<bullet> x = x$i" "x \<bullet> (basis i :: 'a^'n) = (x$i :: 'a::semiring_1)"
-  using i
-  by (auto simp add: dot_def basis_def Cart_lambda_beta cond_application_beta  cond_value_iff setsum_delta cong del: if_weak_cong)
-
-lemma basis_eq_0: "basis i = (0::'a::semiring_1^'n) \<longleftrightarrow> i \<notin> {1..dimindex(UNIV ::'n set)}"
-  by (auto simp add: Cart_eq basis_component zero_index)
+  shows "basis i \<bullet> x = x$i" "x \<bullet> (basis i :: 'a^'n::finite) = (x$i :: 'a::semiring_1)"
+  by (auto simp add: dot_def basis_def cond_application_beta  cond_value_iff setsum_delta cong del: if_weak_cong)
+
+lemma basis_eq_0: "basis i = (0::'a::semiring_1^'n) \<longleftrightarrow> False"
+  by (auto simp add: Cart_eq)
 
 lemma basis_nonzero:
-  assumes k: "k \<in> {1 .. dimindex(UNIV ::'n set)}"
   shows "basis k \<noteq> (0:: 'a::semiring_1 ^'n)"
-  using k by (simp add: basis_eq_0)
-
-lemma vector_eq_ldot: "(\<forall>x. x \<bullet> y = x \<bullet> z) \<longleftrightarrow> y = (z::'a::semiring_1^'n)"
+  by (simp add: basis_eq_0)
+
+lemma vector_eq_ldot: "(\<forall>x. x \<bullet> y = x \<bullet> z) \<longleftrightarrow> y = (z::'a::semiring_1^'n::finite)"
   apply (auto simp add: Cart_eq dot_basis)
   apply (erule_tac x="basis i" in allE)
   apply (simp add: dot_basis)
   apply (subgoal_tac "y = z")
   apply simp
-  apply vector
+  apply (simp add: Cart_eq)
   done
 
-lemma vector_eq_rdot: "(\<forall>z. x \<bullet> z = y \<bullet> z) \<longleftrightarrow> x = (y::'a::semiring_1^'n)"
+lemma vector_eq_rdot: "(\<forall>z. x \<bullet> z = y \<bullet> z) \<longleftrightarrow> x = (y::'a::semiring_1^'n::finite)"
   apply (auto simp add: Cart_eq dot_basis)
   apply (erule_tac x="basis i" in allE)
   apply (simp add: dot_basis)
   apply (subgoal_tac "x = y")
   apply simp
-  apply vector
+  apply (simp add: Cart_eq)
   done
 
 subsection{* Orthogonality. *}
@@ -1294,16 +1290,12 @@
 definition "orthogonal x y \<longleftrightarrow> (x \<bullet> y = 0)"
 
 lemma orthogonal_basis:
-  assumes i:"i \<in> {1 .. dimindex(UNIV ::'n set)}"
-  shows "orthogonal (basis i :: 'a^'n) x \<longleftrightarrow> x$i = (0::'a::ring_1)"
-  using i
-  by (auto simp add: orthogonal_def dot_def basis_def Cart_lambda_beta cond_value_iff cond_application_beta setsum_delta cong del: if_weak_cong)
+  shows "orthogonal (basis i :: 'a^'n::finite) x \<longleftrightarrow> x$i = (0::'a::ring_1)"
+  by (auto simp add: orthogonal_def dot_def basis_def cond_value_iff cond_application_beta setsum_delta cong del: if_weak_cong)
 
 lemma orthogonal_basis_basis:
-  assumes i:"i \<in> {1 .. dimindex(UNIV ::'n set)}"
-  and j: "j \<in> {1 .. dimindex(UNIV ::'n set)}"
-  shows "orthogonal (basis i :: 'a::ring_1^'n) (basis j) \<longleftrightarrow> i \<noteq> j"
-  unfolding orthogonal_basis[OF i] basis_component[OF i] by simp
+  shows "orthogonal (basis i :: 'a::ring_1^'n::finite) (basis j) \<longleftrightarrow> i \<noteq> j"
+  unfolding orthogonal_basis[of i] basis_component[of j] by simp
 
   (* FIXME : Maybe some of these require less than comm_ring, but not all*)
 lemma orthogonal_clauses:
@@ -1326,51 +1318,43 @@
 
 subsection{* Explicit vector construction from lists. *}
 
-lemma Cart_lambda_beta_1[simp]: "(Cart_lambda g)$1 = g 1"
-  apply (rule Cart_lambda_beta[rule_format])
-  using dimindex_ge_1 apply auto done
-
-lemma Cart_lambda_beta_1'[simp]: "(Cart_lambda g)$(Suc 0) = g 1"
-  by (simp only: One_nat_def[symmetric] Cart_lambda_beta_1)
-
-definition "vector l = (\<chi> i. if i <= length l then l ! (i - 1) else 0)"
+primrec from_nat :: "nat \<Rightarrow> 'a::{monoid_add,one}"
+where "from_nat 0 = 0" | "from_nat (Suc n) = 1 + from_nat n"
+
+lemma from_nat [simp]: "from_nat = of_nat"
+by (rule ext, induct_tac x, simp_all)
+
+primrec
+  list_fun :: "nat \<Rightarrow> _ list \<Rightarrow> _ \<Rightarrow> _"
+where
+  "list_fun n [] = (\<lambda>x. 0)"
+| "list_fun n (x # xs) = fun_upd (list_fun (Suc n) xs) (from_nat n) x"
+
+definition "vector l = (\<chi> i. list_fun 1 l i)"
+(*definition "vector l = (\<chi> i. if i <= length l then l ! (i - 1) else 0)"*)
 
 lemma vector_1: "(vector[x]) $1 = x"
-  using dimindex_ge_1
-  by (auto simp add: vector_def Cart_lambda_beta[rule_format])
-lemma dimindex_2[simp]: "2 \<in> {1 .. dimindex (UNIV :: 2 set)}"
-  by (auto simp add: dimindex_def)
-lemma dimindex_2'[simp]: "2 \<in> {Suc 0 .. dimindex (UNIV :: 2 set)}"
-  by (auto simp add: dimindex_def)
-lemma dimindex_3[simp]: "2 \<in> {1 .. dimindex (UNIV :: 3 set)}" "3 \<in> {1 .. dimindex (UNIV :: 3 set)}"
-  by (auto simp add: dimindex_def)
-
-lemma dimindex_3'[simp]: "2 \<in> {Suc 0 .. dimindex (UNIV :: 3 set)}" "3 \<in> {Suc 0 .. dimindex (UNIV :: 3 set)}"
-  by (auto simp add: dimindex_def)
+  unfolding vector_def by simp
 
 lemma vector_2:
  "(vector[x,y]) $1 = x"
  "(vector[x,y] :: 'a^2)$2 = (y::'a::zero)"
-  apply (simp add: vector_def)
-  using Cart_lambda_beta[rule_format, OF dimindex_2, of "\<lambda>i. if i \<le> length [x,y] then [x,y] ! (i - 1) else (0::'a)"]
-  apply (simp only: vector_def )
-  apply auto
-  done
+  unfolding vector_def by simp_all
 
 lemma vector_3:
  "(vector [x,y,z] ::('a::zero)^3)$1 = x"
  "(vector [x,y,z] ::('a::zero)^3)$2 = y"
  "(vector [x,y,z] ::('a::zero)^3)$3 = z"
-apply (simp_all add: vector_def Cart_lambda_beta dimindex_3)
-  using Cart_lambda_beta[rule_format, OF dimindex_3(1), of "\<lambda>i. if i \<le> length [x,y,z] then [x,y,z] ! (i - 1) else (0::'a)"]   using Cart_lambda_beta[rule_format, OF dimindex_3(2), of "\<lambda>i. if i \<le> length [x,y,z] then [x,y,z] ! (i - 1) else (0::'a)"]
-  by simp_all
+  unfolding vector_def by simp_all
 
 lemma forall_vector_1: "(\<forall>v::'a::zero^1. P v) \<longleftrightarrow> (\<forall>x. P(vector[x]))"
   apply auto
   apply (erule_tac x="v$1" in allE)
   apply (subgoal_tac "vector [v$1] = v")
   apply simp
-  by (vector vector_def dimindex_def)
+  apply (vector vector_def)
+  apply (simp add: forall_1)
+  done
 
 lemma forall_vector_2: "(\<forall>v::'a::zero^2. P v) \<longleftrightarrow> (\<forall>x y. P(vector[x, y]))"
   apply auto
@@ -1378,9 +1362,8 @@
   apply (erule_tac x="v$2" in allE)
   apply (subgoal_tac "vector [v$1, v$2] = v")
   apply simp
-  apply (vector vector_def dimindex_def)
-  apply auto
-  apply (subgoal_tac "i = 1 \<or> i =2", auto)
+  apply (vector vector_def)
+  apply (simp add: forall_2)
   done
 
 lemma forall_vector_3: "(\<forall>v::'a::zero^3. P v) \<longleftrightarrow> (\<forall>x y z. P(vector[x, y, z]))"
@@ -1390,9 +1373,8 @@
   apply (erule_tac x="v$3" in allE)
   apply (subgoal_tac "vector [v$1, v$2, v$3] = v")
   apply simp
-  apply (vector vector_def dimindex_def)
-  apply auto
-  apply (subgoal_tac "i = 1 \<or> i =2 \<or> i = 3", auto)
+  apply (vector vector_def)
+  apply (simp add: forall_3)
   done
 
 subsection{* Linear functions. *}
@@ -1400,7 +1382,7 @@
 definition "linear f \<longleftrightarrow> (\<forall>x y. f(x + y) = f x + f y) \<and> (\<forall>c x. f(c *s x) = c *s f x)"
 
 lemma linear_compose_cmul: "linear f ==> linear (\<lambda>x. (c::'a::comm_semiring) *s f x)"
-  by (vector linear_def Cart_eq Cart_lambda_beta[rule_format] ring_simps)
+  by (vector linear_def Cart_eq ring_simps)
 
 lemma linear_compose_neg: "linear (f :: 'a ^'n \<Rightarrow> 'a::comm_ring ^'m) ==> linear (\<lambda>x. -(f(x)))" by (vector linear_def Cart_eq)
 
@@ -1426,9 +1408,9 @@
 
 lemma linear_vmul_component:
   fixes f:: "'a::semiring_1^'m \<Rightarrow> 'a^'n"
-  assumes lf: "linear f" and k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}"
+  assumes lf: "linear f"
   shows "linear (\<lambda>x. f x $ k *s v)"
-  using lf k
+  using lf
   apply (auto simp add: linear_def )
   by (vector ring_simps)+
 
@@ -1485,15 +1467,15 @@
 qed
 
 lemma linear_bounded:
-  fixes f:: "real ^'m \<Rightarrow> real ^'n"
+  fixes f:: "real ^'m::finite \<Rightarrow> real ^'n::finite"
   assumes lf: "linear f"
   shows "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x"
 proof-
-  let ?S = "{1..dimindex(UNIV:: 'm set)}"
+  let ?S = "UNIV:: 'm set"
   let ?B = "setsum (\<lambda>i. norm(f(basis i))) ?S"
   have fS: "finite ?S" by simp
   {fix x:: "real ^ 'm"
-    let ?g = "(\<lambda>i::nat. (x$i) *s (basis i) :: real ^ 'm)"
+    let ?g = "(\<lambda>i. (x$i) *s (basis i) :: real ^ 'm)"
     have "norm (f x) = norm (f (setsum (\<lambda>i. (x$i) *s (basis i)) ?S))"
       by (simp only:  basis_expansion)
     also have "\<dots> = norm (setsum (\<lambda>i. (x$i) *s f (basis i))?S)"
@@ -1501,7 +1483,7 @@
       by auto
     finally have th0: "norm (f x) = norm (setsum (\<lambda>i. (x$i) *s f (basis i))?S)" .
     {fix i assume i: "i \<in> ?S"
-      from component_le_norm[OF i, of x]
+      from component_le_norm[of x i]
       have "norm ((x$i) *s f (basis i :: real ^'m)) \<le> norm (f (basis i)) * norm x"
       unfolding norm_mul
       apply (simp only: mult_commute)
@@ -1514,7 +1496,7 @@
 qed
 
 lemma linear_bounded_pos:
-  fixes f:: "real ^'n \<Rightarrow> real ^ 'm"
+  fixes f:: "real ^'n::finite \<Rightarrow> real ^ 'm::finite"
   assumes lf: "linear f"
   shows "\<exists>B > 0. \<forall>x. norm (f x) \<le> B * norm x"
 proof-
@@ -1595,12 +1577,12 @@
 qed
 
 lemma bilinear_bounded:
-  fixes h:: "real ^'m \<Rightarrow> real^'n \<Rightarrow> real ^ 'k"
+  fixes h:: "real ^'m::finite \<Rightarrow> real^'n::finite \<Rightarrow> real ^ 'k::finite"
   assumes bh: "bilinear h"
   shows "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"
 proof-
-  let ?M = "{1 .. dimindex (UNIV :: 'm set)}"
-  let ?N = "{1 .. dimindex (UNIV :: 'n set)}"
+  let ?M = "UNIV :: 'm set"
+  let ?N = "UNIV :: 'n set"
   let ?B = "setsum (\<lambda>(i,j). norm (h (basis i) (basis j))) (?M \<times> ?N)"
   have fM: "finite ?M" and fN: "finite ?N" by simp_all
   {fix x:: "real ^ 'm" and  y :: "real^'n"
@@ -1622,7 +1604,7 @@
 qed
 
 lemma bilinear_bounded_pos:
-  fixes h:: "real ^'m \<Rightarrow> real^'n \<Rightarrow> real ^ 'k"
+  fixes h:: "real ^'m::finite \<Rightarrow> real^'n::finite \<Rightarrow> real ^ 'k::finite"
   assumes bh: "bilinear h"
   shows "\<exists>B > 0. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"
 proof-
@@ -1649,12 +1631,12 @@
 lemma choice_iff: "(\<forall>x. \<exists>y. P x y) \<longleftrightarrow> (\<exists>f. \<forall>x. P x (f x))" by metis
 
 lemma adjoint_works_lemma:
-  fixes f:: "'a::ring_1 ^'n \<Rightarrow> 'a ^ 'm"
+  fixes f:: "'a::ring_1 ^'n::finite \<Rightarrow> 'a ^ 'm::finite"
   assumes lf: "linear f"
   shows "\<forall>x y. f x \<bullet> y = x \<bullet> adjoint f y"
 proof-
-  let ?N = "{1 .. dimindex (UNIV :: 'n set)}"
-  let ?M = "{1 .. dimindex (UNIV :: 'm set)}"
+  let ?N = "UNIV :: 'n set"
+  let ?M = "UNIV :: 'm set"
   have fN: "finite ?N" by simp
   have fM: "finite ?M" by simp
   {fix y:: "'a ^ 'm"
@@ -1667,7 +1649,7 @@
 	by (simp add: linear_cmul[OF lf])
       finally have "f x \<bullet> y = x \<bullet> ?w"
 	apply (simp only: )
-	apply (simp add: dot_def setsum_component Cart_lambda_beta setsum_left_distrib setsum_right_distrib vector_component setsum_commute[of _ ?M ?N] ring_simps del: One_nat_def)
+	apply (simp add: dot_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] ring_simps)
 	done}
   }
   then show ?thesis unfolding adjoint_def
@@ -1677,34 +1659,34 @@
 qed
 
 lemma adjoint_works:
-  fixes f:: "'a::ring_1 ^'n \<Rightarrow> 'a ^ 'm"
+  fixes f:: "'a::ring_1 ^'n::finite \<Rightarrow> 'a ^ 'm::finite"
   assumes lf: "linear f"
   shows "x \<bullet> adjoint f y = f x \<bullet> y"
   using adjoint_works_lemma[OF lf] by metis
 
 
 lemma adjoint_linear:
-  fixes f :: "'a::comm_ring_1 ^'n \<Rightarrow> 'a ^ 'm"
+  fixes f :: "'a::comm_ring_1 ^'n::finite \<Rightarrow> 'a ^ 'm::finite"
   assumes lf: "linear f"
   shows "linear (adjoint f)"
   by (simp add: linear_def vector_eq_ldot[symmetric] dot_radd dot_rmult adjoint_works[OF lf])
 
 lemma adjoint_clauses:
-  fixes f:: "'a::comm_ring_1 ^'n \<Rightarrow> 'a ^ 'm"
+  fixes f:: "'a::comm_ring_1 ^'n::finite \<Rightarrow> 'a ^ 'm::finite"
   assumes lf: "linear f"
   shows "x \<bullet> adjoint f y = f x \<bullet> y"
   and "adjoint f y \<bullet> x = y \<bullet> f x"
   by (simp_all add: adjoint_works[OF lf] dot_sym )
 
 lemma adjoint_adjoint:
-  fixes f:: "'a::comm_ring_1 ^ 'n \<Rightarrow> _"
+  fixes f:: "'a::comm_ring_1 ^ 'n::finite \<Rightarrow> 'a ^ 'm::finite"
   assumes lf: "linear f"
   shows "adjoint (adjoint f) = f"
   apply (rule ext)
   by (simp add: vector_eq_ldot[symmetric] adjoint_clauses[OF adjoint_linear[OF lf]] adjoint_clauses[OF lf])
 
 lemma adjoint_unique:
-  fixes f:: "'a::comm_ring_1 ^ 'n \<Rightarrow> 'a ^ 'm"
+  fixes f:: "'a::comm_ring_1 ^ 'n::finite \<Rightarrow> 'a ^ 'm::finite"
   assumes lf: "linear f" and u: "\<forall>x y. f' x \<bullet> y = x \<bullet> f y"
   shows "f' = adjoint f"
   apply (rule ext)
@@ -1716,14 +1698,14 @@
 consts generic_mult :: "'a \<Rightarrow> 'b \<Rightarrow> 'c" (infixr "\<star>" 75)
 
 defs (overloaded)
-matrix_matrix_mult_def: "(m:: ('a::semiring_1) ^'n^'m) \<star> (m' :: 'a ^'p^'n) \<equiv> (\<chi> i j. setsum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) {1 .. dimindex (UNIV :: 'n set)}) ::'a ^ 'p ^'m"
+matrix_matrix_mult_def: "(m:: ('a::semiring_1) ^'n^'m) \<star> (m' :: 'a ^'p^'n) \<equiv> (\<chi> i j. setsum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m"
 
 abbreviation
   matrix_matrix_mult' :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"  (infixl "**" 70)
   where "m ** m' == m\<star> m'"
 
 defs (overloaded)
-  matrix_vector_mult_def: "(m::('a::semiring_1) ^'n^'m) \<star> (x::'a ^'n) \<equiv> (\<chi> i. setsum (\<lambda>j. ((m$i)$j) * (x$j)) {1..dimindex(UNIV ::'n set)}) :: 'a^'m"
+  matrix_vector_mult_def: "(m::('a::semiring_1) ^'n^'m) \<star> (x::'a ^'n) \<equiv> (\<chi> i. setsum (\<lambda>j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m"
 
 abbreviation
   matrix_vector_mult' :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"  (infixl "*v" 70)
@@ -1731,19 +1713,19 @@
   "m *v v == m \<star> v"
 
 defs (overloaded)
-  vector_matrix_mult_def: "(x::'a^'m) \<star> (m::('a::semiring_1) ^'n^'m) \<equiv> (\<chi> j. setsum (\<lambda>i. ((m$i)$j) * (x$i)) {1..dimindex(UNIV :: 'm set)}) :: 'a^'n"
+  vector_matrix_mult_def: "(x::'a^'m) \<star> (m::('a::semiring_1) ^'n^'m) \<equiv> (\<chi> j. setsum (\<lambda>i. ((m$i)$j) * (x$i)) (UNIV :: 'm set)) :: 'a^'n"
 
 abbreviation
   vactor_matrix_mult' :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "  (infixl "v*" 70)
   where
   "v v* m == v \<star> m"
 
-definition "(mat::'a::zero => 'a ^'n^'m) k = (\<chi> i j. if i = j then k else 0)"
+definition "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)"
 definition "(transp::'a^'n^'m \<Rightarrow> 'a^'m^'n) A = (\<chi> i j. ((A$j)$i))"
-definition "(row::nat => 'a ^'n^'m \<Rightarrow> 'a ^'n) i A = (\<chi> j. ((A$i)$j))"
-definition "(column::nat =>'a^'n^'m =>'a^'m) j A = (\<chi> i. ((A$i)$j))"
-definition "rows(A::'a^'n^'m) = { row i A | i. i \<in> {1 .. dimindex(UNIV :: 'm set)}}"
-definition "columns(A::'a^'n^'m) = { column i A | i. i \<in> {1 .. dimindex(UNIV :: 'n set)}}"
+definition "(row::'m => 'a ^'n^'m \<Rightarrow> 'a ^'n) i A = (\<chi> j. ((A$i)$j))"
+definition "(column::'n =>'a^'n^'m =>'a^'m) j A = (\<chi> i. ((A$i)$j))"
+definition "rows(A::'a^'n^'m) = { row i A | i. i \<in> (UNIV :: 'm set)}"
+definition "columns(A::'a^'n^'m) = { column i A | i. i \<in> (UNIV :: 'n set)}"
 
 lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def)
 lemma matrix_add_ldistrib: "(A ** (B + C)) = (A \<star> B) + (A \<star> C)"
@@ -1756,16 +1738,20 @@
   using setsum_delta[OF fS, of a b, symmetric]
   by (auto intro: setsum_cong)
 
-lemma matrix_mul_lid: "mat 1 ** A = A"
+lemma matrix_mul_lid:
+  fixes A :: "'a::semiring_1 ^ 'm ^ 'n::finite"
+  shows "mat 1 ** A = A"
   apply (simp add: matrix_matrix_mult_def mat_def)
   apply vector
-  by (auto simp only: cond_value_iff cond_application_beta setsum_delta'[OF finite_atLeastAtMost]  mult_1_left mult_zero_left if_True)
-
-
-lemma matrix_mul_rid: "A ** mat 1 = A"
+  by (auto simp only: cond_value_iff cond_application_beta setsum_delta'[OF finite]  mult_1_left mult_zero_left if_True UNIV_I)
+
+
+lemma matrix_mul_rid:
+  fixes A :: "'a::semiring_1 ^ 'm::finite ^ 'n"
+  shows "A ** mat 1 = A"
   apply (simp add: matrix_matrix_mult_def mat_def)
   apply vector
-  by (auto simp only: cond_value_iff cond_application_beta setsum_delta[OF finite_atLeastAtMost]  mult_1_right mult_zero_right if_True cong: if_cong)
+  by (auto simp only: cond_value_iff cond_application_beta setsum_delta[OF finite]  mult_1_right mult_zero_right if_True UNIV_I cong: if_cong)
 
 lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C"
   apply (vector matrix_matrix_mult_def setsum_right_distrib setsum_left_distrib mult_assoc)
@@ -1779,31 +1765,31 @@
   apply simp
   done
 
-lemma matrix_vector_mul_lid: "mat 1 *v x = x"
+lemma matrix_vector_mul_lid: "mat 1 *v x = (x::'a::semiring_1 ^ 'n::finite)"
   apply (vector matrix_vector_mult_def mat_def)
   by (simp add: cond_value_iff cond_application_beta
     setsum_delta' cong del: if_weak_cong)
 
 lemma matrix_transp_mul: "transp(A ** B) = transp B ** transp (A::'a::comm_semiring_1^'m^'n)"
-  by (simp add: matrix_matrix_mult_def transp_def Cart_eq Cart_lambda_beta mult_commute)
-
-lemma matrix_eq: "A = B \<longleftrightarrow>  (\<forall>x. A *v x = B *v x)" (is "?lhs \<longleftrightarrow> ?rhs")
+  by (simp add: matrix_matrix_mult_def transp_def Cart_eq mult_commute)
+
+lemma matrix_eq:
+  fixes A B :: "'a::semiring_1 ^ 'n::finite ^ 'm"
+  shows "A = B \<longleftrightarrow>  (\<forall>x. A *v x = B *v x)" (is "?lhs \<longleftrightarrow> ?rhs")
   apply auto
   apply (subst Cart_eq)
   apply clarify
-  apply (clarsimp simp add: matrix_vector_mult_def basis_def cond_value_iff cond_application_beta Cart_eq Cart_lambda_beta cong del: if_weak_cong)
+  apply (clarsimp simp add: matrix_vector_mult_def basis_def cond_value_iff cond_application_beta Cart_eq cong del: if_weak_cong)
   apply (erule_tac x="basis ia" in allE)
-  apply (erule_tac x="i" in ballE)
-  by (auto simp add: basis_def cond_value_iff cond_application_beta Cart_lambda_beta setsum_delta[OF finite_atLeastAtMost] cong del: if_weak_cong)
+  apply (erule_tac x="i" in allE)
+  by (auto simp add: basis_def cond_value_iff cond_application_beta setsum_delta[OF finite] cong del: if_weak_cong)
 
 lemma matrix_vector_mul_component:
-  assumes k: "k \<in> {1.. dimindex (UNIV :: 'm set)}"
   shows "((A::'a::semiring_1^'n'^'m) *v x)$k = (A$k) \<bullet> x"
-  using k
-  by (simp add: matrix_vector_mult_def Cart_lambda_beta dot_def)
+  by (simp add: matrix_vector_mult_def dot_def)
 
 lemma dot_lmul_matrix: "((x::'a::comm_semiring_1 ^'n) v* A) \<bullet> y = x \<bullet> (A *v y)"
-  apply (simp add: dot_def matrix_vector_mult_def vector_matrix_mult_def setsum_left_distrib setsum_right_distrib Cart_lambda_beta mult_ac)
+  apply (simp add: dot_def matrix_vector_mult_def vector_matrix_mult_def setsum_left_distrib setsum_right_distrib mult_ac)
   apply (subst setsum_commute)
   by simp
 
@@ -1815,23 +1801,16 @@
 
 lemma row_transp:
   fixes A:: "'a::semiring_1^'n^'m"
-  assumes i: "i \<in> {1.. dimindex (UNIV :: 'n set)}"
   shows "row i (transp A) = column i A"
-  using i
-  by (simp add: row_def column_def transp_def Cart_eq Cart_lambda_beta)
+  by (simp add: row_def column_def transp_def Cart_eq)
 
 lemma column_transp:
   fixes A:: "'a::semiring_1^'n^'m"
-  assumes i: "i \<in> {1.. dimindex (UNIV :: 'm set)}"
   shows "column i (transp A) = row i A"
-  using i
-  by (simp add: row_def column_def transp_def Cart_eq Cart_lambda_beta)
+  by (simp add: row_def column_def transp_def Cart_eq)
 
 lemma rows_transp: "rows(transp (A::'a::semiring_1^'n^'m)) = columns A"
-apply (auto simp add: rows_def columns_def row_transp intro: set_ext)
-apply (rule_tac x=i in exI)
-apply (auto simp add: row_transp)
-done
+by (auto simp add: rows_def columns_def row_transp intro: set_ext)
 
 lemma columns_transp: "columns(transp (A::'a::semiring_1^'n^'m)) = rows A" by (metis transp_transp rows_transp)
 
@@ -1840,25 +1819,25 @@
 lemma matrix_mult_dot: "A *v x = (\<chi> i. A$i \<bullet> x)"
   by (simp add: matrix_vector_mult_def dot_def)
 
-lemma matrix_mult_vsum: "(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\<lambda>i. (x$i) *s column i A) {1 .. dimindex(UNIV:: 'n set)}"
-  by (simp add: matrix_vector_mult_def Cart_eq setsum_component Cart_lambda_beta vector_component column_def mult_commute)
+lemma matrix_mult_vsum: "(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\<lambda>i. (x$i) *s column i A) (UNIV:: 'n set)"
+  by (simp add: matrix_vector_mult_def Cart_eq column_def mult_commute)
 
 lemma vector_componentwise:
-  "(x::'a::ring_1^'n) = (\<chi> j. setsum (\<lambda>i. (x$i) * (basis i :: 'a^'n)$j) {1..dimindex(UNIV :: 'n set)})"
+  "(x::'a::ring_1^'n::finite) = (\<chi> j. setsum (\<lambda>i. (x$i) * (basis i :: 'a^'n)$j) (UNIV :: 'n set))"
   apply (subst basis_expansion[symmetric])
-  by (vector Cart_eq Cart_lambda_beta setsum_component)
+  by (vector Cart_eq setsum_component)
 
 lemma linear_componentwise:
-  fixes f:: "'a::ring_1 ^ 'm \<Rightarrow> 'a ^ 'n"
-  assumes lf: "linear f" and j: "j \<in> {1 .. dimindex (UNIV :: 'n set)}"
-  shows "(f x)$j = setsum (\<lambda>i. (x$i) * (f (basis i)$j)) {1 .. dimindex (UNIV :: 'm set)}" (is "?lhs = ?rhs")
+  fixes f:: "'a::ring_1 ^ 'm::finite \<Rightarrow> 'a ^ 'n"
+  assumes lf: "linear f"
+  shows "(f x)$j = setsum (\<lambda>i. (x$i) * (f (basis i)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs")
 proof-
-  let ?M = "{1 .. dimindex (UNIV :: 'm set)}"
-  let ?N = "{1 .. dimindex (UNIV :: 'n set)}"
+  let ?M = "(UNIV :: 'm set)"
+  let ?N = "(UNIV :: 'n set)"
   have fM: "finite ?M" by simp
   have "?rhs = (setsum (\<lambda>i.(x$i) *s f (basis i) ) ?M)$j"
-    unfolding vector_smult_component[OF j, symmetric]
-    unfolding setsum_component[OF j, of "(\<lambda>i.(x$i) *s f (basis i :: 'a^'m))" ?M]
+    unfolding vector_smult_component[symmetric]
+    unfolding setsum_component[of "(\<lambda>i.(x$i) *s f (basis i :: 'a^'m))" ?M]
     ..
   then show ?thesis unfolding linear_setsum_mul[OF lf fM, symmetric] basis_expansion ..
 qed
@@ -1876,38 +1855,38 @@
 where "matrix f = (\<chi> i j. (f(basis j))$i)"
 
 lemma matrix_vector_mul_linear: "linear(\<lambda>x. A *v (x::'a::comm_semiring_1 ^ 'n))"
-  by (simp add: linear_def matrix_vector_mult_def Cart_eq Cart_lambda_beta vector_component ring_simps setsum_right_distrib setsum_addf)
-
-lemma matrix_works: assumes lf: "linear f" shows "matrix f *v x = f (x::'a::comm_ring_1 ^ 'n)"
-apply (simp add: matrix_def matrix_vector_mult_def Cart_eq Cart_lambda_beta mult_commute del: One_nat_def)
+  by (simp add: linear_def matrix_vector_mult_def Cart_eq ring_simps setsum_right_distrib setsum_addf)
+
+lemma matrix_works: assumes lf: "linear f" shows "matrix f *v x = f (x::'a::comm_ring_1 ^ 'n::finite)"
+apply (simp add: matrix_def matrix_vector_mult_def Cart_eq mult_commute)
 apply clarify
 apply (rule linear_componentwise[OF lf, symmetric])
-apply simp
 done
 
-lemma matrix_vector_mul: "linear f ==> f = (\<lambda>x. matrix f *v (x::'a::comm_ring_1 ^ 'n))" by (simp add: ext matrix_works)
-
-lemma matrix_of_matrix_vector_mul: "matrix(\<lambda>x. A *v (x :: 'a:: comm_ring_1 ^ 'n)) = A"
+lemma matrix_vector_mul: "linear f ==> f = (\<lambda>x. matrix f *v (x::'a::comm_ring_1 ^ 'n::finite))" by (simp add: ext matrix_works)
+
+lemma matrix_of_matrix_vector_mul: "matrix(\<lambda>x. A *v (x :: 'a:: comm_ring_1 ^ 'n::finite)) = A"
   by (simp add: matrix_eq matrix_vector_mul_linear matrix_works)
 
 lemma matrix_compose:
-  assumes lf: "linear (f::'a::comm_ring_1^'n \<Rightarrow> _)" and lg: "linear g"
+  assumes lf: "linear (f::'a::comm_ring_1^'n::finite \<Rightarrow> 'a^'m::finite)"
+  and lg: "linear (g::'a::comm_ring_1^'m::finite \<Rightarrow> 'a^'k)"
   shows "matrix (g o f) = matrix g ** matrix f"
   using lf lg linear_compose[OF lf lg] matrix_works[OF linear_compose[OF lf lg]]
   by (simp  add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def)
 
-lemma matrix_vector_column:"(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\<lambda>i. (x$i) *s ((transp A)$i)) {1..dimindex(UNIV:: 'n set)}"
-  by (simp add: matrix_vector_mult_def transp_def Cart_eq Cart_lambda_beta setsum_component vector_component mult_commute)
-
-lemma adjoint_matrix: "adjoint(\<lambda>x. (A::'a::comm_ring_1^'n^'m) *v x) = (\<lambda>x. transp A *v x)"
+lemma matrix_vector_column:"(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\<lambda>i. (x$i) *s ((transp A)$i)) (UNIV:: 'n set)"
+  by (simp add: matrix_vector_mult_def transp_def Cart_eq mult_commute)
+
+lemma adjoint_matrix: "adjoint(\<lambda>x. (A::'a::comm_ring_1^'n::finite^'m::finite) *v x) = (\<lambda>x. transp A *v x)"
   apply (rule adjoint_unique[symmetric])
   apply (rule matrix_vector_mul_linear)
-  apply (simp add: transp_def dot_def Cart_lambda_beta matrix_vector_mult_def setsum_left_distrib setsum_right_distrib)
+  apply (simp add: transp_def dot_def matrix_vector_mult_def setsum_left_distrib setsum_right_distrib)
   apply (subst setsum_commute)
   apply (auto simp add: mult_ac)
   done
 
-lemma matrix_adjoint: assumes lf: "linear (f :: 'a::comm_ring_1^'n \<Rightarrow> 'a ^ 'm)"
+lemma matrix_adjoint: assumes lf: "linear (f :: 'a::comm_ring_1^'n::finite \<Rightarrow> 'a ^ 'm::finite)"
   shows "matrix(adjoint f) = transp(matrix f)"
   apply (subst matrix_vector_mul[OF lf])
   unfolding adjoint_matrix matrix_of_matrix_vector_mul ..
@@ -1980,21 +1959,21 @@
 qed
 
 
-lemma lambda_skolem: "(\<forall>i \<in> {1 .. dimindex(UNIV :: 'n set)}. \<exists>x. P i x) \<longleftrightarrow>
-   (\<exists>x::'a ^ 'n. \<forall>i \<in> {1 .. dimindex(UNIV:: 'n set)}. P i (x$i))" (is "?lhs \<longleftrightarrow> ?rhs")
+lemma lambda_skolem: "(\<forall>i. \<exists>x. P i x) \<longleftrightarrow>
+   (\<exists>x::'a ^ 'n. \<forall>i. P i (x$i))" (is "?lhs \<longleftrightarrow> ?rhs")
 proof-
-  let ?S = "{1 .. dimindex(UNIV :: 'n set)}"
+  let ?S = "(UNIV :: 'n set)"
   {assume H: "?rhs"
     then have ?lhs by auto}
   moreover
   {assume H: "?lhs"
-    then obtain f where f:"\<forall>i\<in> ?S. P i (f i)" unfolding Ball_def choice_iff by metis
+    then obtain f where f:"\<forall>i. P i (f i)" unfolding choice_iff by metis
     let ?x = "(\<chi> i. (f i)) :: 'a ^ 'n"
-    {fix i assume i: "i \<in> ?S"
-      with f i have "P i (f i)" by metis
-      then have "P i (?x$i)" using Cart_lambda_beta[of f, rule_format, OF i] by auto
+    {fix i
+      from f have "P i (f i)" by metis
+      then have "P i (?x$i)" by auto
     }
-    hence "\<forall>i \<in> ?S. P i (?x$i)" by metis
+    hence "\<forall>i. P i (?x$i)" by metis
     hence ?rhs by metis }
   ultimately show ?thesis by metis
 qed
@@ -2237,7 +2216,7 @@
 definition "onorm f = rsup {norm (f x)| x. norm x = 1}"
 
 lemma norm_bound_generalize:
-  fixes f:: "real ^'n \<Rightarrow> real^'m"
+  fixes f:: "real ^'n::finite \<Rightarrow> real^'m::finite"
   assumes lf: "linear f"
   shows "(\<forall>x. norm x = 1 \<longrightarrow> norm (f x) \<le> b) \<longleftrightarrow> (\<forall>x. norm (f x) \<le> b * norm x)" (is "?lhs \<longleftrightarrow> ?rhs")
 proof-
@@ -2248,8 +2227,8 @@
 
   moreover
   {assume H: ?lhs
-    from H[rule_format, of "basis 1"]
-    have bp: "b \<ge> 0" using norm_ge_zero[of "f (basis 1)"] dimindex_ge_1[of "UNIV:: 'n set"]
+    from H[rule_format, of "basis arbitrary"]
+    have bp: "b \<ge> 0" using norm_ge_zero[of "f (basis arbitrary)"]
       by (auto simp add: norm_basis elim: order_trans [OF norm_ge_zero])
     {fix x :: "real ^'n"
       {assume "x = 0"
@@ -2270,14 +2249,14 @@
 qed
 
 lemma onorm:
-  fixes f:: "real ^'n \<Rightarrow> real ^'m"
+  fixes f:: "real ^'n::finite \<Rightarrow> real ^'m::finite"
   assumes lf: "linear f"
   shows "norm (f x) <= onorm f * norm x"
   and "\<forall>x. norm (f x) <= b * norm x \<Longrightarrow> onorm f <= b"
 proof-
   {
     let ?S = "{norm (f x) |x. norm x = 1}"
-    have Se: "?S \<noteq> {}" using  norm_basis_1 by auto
+    have Se: "?S \<noteq> {}" using  norm_basis by auto
     from linear_bounded[OF lf] have b: "\<exists> b. ?S *<= b"
       unfolding norm_bound_generalize[OF lf, symmetric] by (auto simp add: setle_def)
     {from rsup[OF Se b, unfolded onorm_def[symmetric]]
@@ -2294,10 +2273,10 @@
   }
 qed
 
-lemma onorm_pos_le: assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" shows "0 <= onorm f"
-  using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "basis 1"], unfolded norm_basis_1] by simp
-
-lemma onorm_eq_0: assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)"
+lemma onorm_pos_le: assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite)" shows "0 <= onorm f"
+  using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "basis arbitrary"], unfolded norm_basis] by simp
+
+lemma onorm_eq_0: assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite)"
   shows "onorm f = 0 \<longleftrightarrow> (\<forall>x. f x = 0)"
   using onorm[OF lf]
   apply (auto simp add: onorm_pos_le)
@@ -2307,7 +2286,7 @@
   apply arith
   done
 
-lemma onorm_const: "onorm(\<lambda>x::real^'n. (y::real ^ 'm)) = norm y"
+lemma onorm_const: "onorm(\<lambda>x::real^'n::finite. (y::real ^ 'm::finite)) = norm y"
 proof-
   let ?f = "\<lambda>x::real^'n. (y::real ^ 'm)"
   have th: "{norm (?f x)| x. norm x = 1} = {norm y}"
@@ -2317,13 +2296,14 @@
     apply (rule rsup_unique) by (simp_all  add: setle_def)
 qed
 
-lemma onorm_pos_lt: assumes lf: "linear (f::real ^ 'n \<Rightarrow> real ^'m)"
+lemma onorm_pos_lt: assumes lf: "linear (f::real ^ 'n::finite \<Rightarrow> real ^'m::finite)"
   shows "0 < onorm f \<longleftrightarrow> ~(\<forall>x. f x = 0)"
   unfolding onorm_eq_0[OF lf, symmetric]
   using onorm_pos_le[OF lf] by arith
 
 lemma onorm_compose:
-  assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" and lg: "linear g"
+  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite)"
+  and lg: "linear (g::real^'k::finite \<Rightarrow> real^'n::finite)"
   shows "onorm (f o g) <= onorm f * onorm g"
   apply (rule onorm(2)[OF linear_compose[OF lg lf], rule_format])
   unfolding o_def
@@ -2335,18 +2315,18 @@
   apply (rule onorm_pos_le[OF lf])
   done
 
-lemma onorm_neg_lemma: assumes lf: "linear (f::real ^'n \<Rightarrow> real^'m)"
+lemma onorm_neg_lemma: assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real^'m::finite)"
   shows "onorm (\<lambda>x. - f x) \<le> onorm f"
   using onorm[OF linear_compose_neg[OF lf]] onorm[OF lf]
   unfolding norm_minus_cancel by metis
 
-lemma onorm_neg: assumes lf: "linear (f::real ^'n \<Rightarrow> real^'m)"
+lemma onorm_neg: assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real^'m::finite)"
   shows "onorm (\<lambda>x. - f x) = onorm f"
   using onorm_neg_lemma[OF lf] onorm_neg_lemma[OF linear_compose_neg[OF lf]]
   by simp
 
 lemma onorm_triangle:
-  assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" and lg: "linear g"
+  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite)" and lg: "linear g"
   shows "onorm (\<lambda>x. f x + g x) <= onorm f + onorm g"
   apply(rule onorm(2)[OF linear_compose_add[OF lf lg], rule_format])
   apply (rule order_trans)
@@ -2357,14 +2337,14 @@
   apply (rule onorm(1)[OF lg])
   done
 
-lemma onorm_triangle_le: "linear (f::real ^'n \<Rightarrow> real ^'m) \<Longrightarrow> linear g \<Longrightarrow> onorm(f) + onorm(g) <= e
+lemma onorm_triangle_le: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite) \<Longrightarrow> linear g \<Longrightarrow> onorm(f) + onorm(g) <= e
   \<Longrightarrow> onorm(\<lambda>x. f x + g x) <= e"
   apply (rule order_trans)
   apply (rule onorm_triangle)
   apply assumption+
   done
 
-lemma onorm_triangle_lt: "linear (f::real ^'n \<Rightarrow> real ^'m) \<Longrightarrow> linear g \<Longrightarrow> onorm(f) + onorm(g) < e
+lemma onorm_triangle_lt: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite) \<Longrightarrow> linear g \<Longrightarrow> onorm(f) + onorm(g) < e
   ==> onorm(\<lambda>x. f x + g x) < e"
   apply (rule order_le_less_trans)
   apply (rule onorm_triangle)
@@ -2381,7 +2361,7 @@
   by (simp add: vec1_def)
 
 lemma vec1_dest_vec1[simp]: "vec1(dest_vec1 x) = x" "dest_vec1(vec1 y) = y"
-  by (simp_all add: vec1_def dest_vec1_def Cart_eq Cart_lambda_beta dimindex_def del: One_nat_def)
+  by (simp_all add: vec1_def dest_vec1_def Cart_eq forall_1)
 
 lemma forall_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (vec1 x))" by (metis vec1_dest_vec1)
 
@@ -2451,21 +2431,21 @@
   shows "linear f \<Longrightarrow> linear (\<lambda>x. dest_vec1(f x) *s v)"
   unfolding dest_vec1_def
   apply (rule linear_vmul_component)
-  by (auto simp add: dimindex_def)
+  by auto
 
 lemma linear_from_scalars:
   assumes lf: "linear (f::'a::comm_ring_1 ^1 \<Rightarrow> 'a^'n)"
   shows "f = (\<lambda>x. dest_vec1 x *s column 1 (matrix f))"
   apply (rule ext)
   apply (subst matrix_works[OF lf, symmetric])
-  apply (auto simp add: Cart_eq matrix_vector_mult_def dest_vec1_def column_def Cart_lambda_beta vector_component dimindex_def mult_commute del: One_nat_def )
+  apply (auto simp add: Cart_eq matrix_vector_mult_def dest_vec1_def column_def  mult_commute UNIV_1)
   done
 
-lemma linear_to_scalars: assumes lf: "linear (f::'a::comm_ring_1 ^'n \<Rightarrow> 'a^1)"
+lemma linear_to_scalars: assumes lf: "linear (f::'a::comm_ring_1 ^'n::finite \<Rightarrow> 'a^1)"
   shows "f = (\<lambda>x. vec1(row 1 (matrix f) \<bullet> x))"
   apply (rule ext)
   apply (subst matrix_works[OF lf, symmetric])
-  apply (auto simp add: Cart_eq matrix_vector_mult_def vec1_def row_def Cart_lambda_beta vector_component dimindex_def dot_def mult_commute)
+  apply (simp add: Cart_eq matrix_vector_mult_def vec1_def row_def dot_def mult_commute forall_1)
   done
 
 lemma dest_vec1_eq_0: "dest_vec1 x = 0 \<longleftrightarrow> x = 0"
@@ -2485,25 +2465,25 @@
 text{* Pasting vectors. *}
 
 lemma linear_fstcart: "linear fstcart"
-  by (auto simp add: linear_def fstcart_def Cart_eq Cart_lambda_beta vector_component dimindex_finite_sum)
+  by (auto simp add: linear_def Cart_eq)
 
 lemma linear_sndcart: "linear sndcart"
-  by (auto simp add: linear_def sndcart_def Cart_eq Cart_lambda_beta vector_component dimindex_finite_sum)
+  by (auto simp add: linear_def Cart_eq)
 
 lemma fstcart_vec[simp]: "fstcart(vec x) = vec x"
-  by (vector fstcart_def vec_def dimindex_finite_sum)
-
-lemma fstcart_add[simp]:"fstcart(x + y) = fstcart (x::'a::{plus,times}^('b,'c) finite_sum) + fstcart y"
-  using linear_fstcart[unfolded linear_def] by blast
-
-lemma fstcart_cmul[simp]:"fstcart(c*s x) = c*s fstcart (x::'a::{plus,times}^('b,'c) finite_sum)"
-  using linear_fstcart[unfolded linear_def] by blast
-
-lemma fstcart_neg[simp]:"fstcart(- x) = - fstcart (x::'a::ring_1^('b,'c) finite_sum)"
-unfolding vector_sneg_minus1 fstcart_cmul ..
-
-lemma fstcart_sub[simp]:"fstcart(x - y) = fstcart (x::'a::ring_1^('b,'c) finite_sum) - fstcart y"
-  unfolding diff_def fstcart_add fstcart_neg  ..
+  by (simp add: Cart_eq)
+
+lemma fstcart_add[simp]:"fstcart(x + y) = fstcart (x::'a::{plus,times}^('b + 'c)) + fstcart y"
+  by (simp add: Cart_eq)
+
+lemma fstcart_cmul[simp]:"fstcart(c*s x) = c*s fstcart (x::'a::{plus,times}^('b + 'c))"
+  by (simp add: Cart_eq)
+
+lemma fstcart_neg[simp]:"fstcart(- x) = - fstcart (x::'a::ring_1^('b + 'c))"
+  by (simp add: Cart_eq)
+
+lemma fstcart_sub[simp]:"fstcart(x - y) = fstcart (x::'a::ring_1^('b + 'c)) - fstcart y"
+  by (simp add: Cart_eq)
 
 lemma fstcart_setsum:
   fixes f:: "'d \<Rightarrow> 'a::semiring_1^_"
@@ -2512,19 +2492,19 @@
   by (induct rule: finite_induct[OF fS], simp_all add: vec_0[symmetric] del: vec_0)
 
 lemma sndcart_vec[simp]: "sndcart(vec x) = vec x"
-  by (vector sndcart_def vec_def dimindex_finite_sum)
-
-lemma sndcart_add[simp]:"sndcart(x + y) = sndcart (x::'a::{plus,times}^('b,'c) finite_sum) + sndcart y"
-  using linear_sndcart[unfolded linear_def] by blast
-
-lemma sndcart_cmul[simp]:"sndcart(c*s x) = c*s sndcart (x::'a::{plus,times}^('b,'c) finite_sum)"
-  using linear_sndcart[unfolded linear_def] by blast
-
-lemma sndcart_neg[simp]:"sndcart(- x) = - sndcart (x::'a::ring_1^('b,'c) finite_sum)"
-unfolding vector_sneg_minus1 sndcart_cmul ..
-
-lemma sndcart_sub[simp]:"sndcart(x - y) = sndcart (x::'a::ring_1^('b,'c) finite_sum) - sndcart y"
-  unfolding diff_def sndcart_add sndcart_neg  ..
+  by (simp add: Cart_eq)
+
+lemma sndcart_add[simp]:"sndcart(x + y) = sndcart (x::'a::{plus,times}^('b + 'c)) + sndcart y"
+  by (simp add: Cart_eq)
+
+lemma sndcart_cmul[simp]:"sndcart(c*s x) = c*s sndcart (x::'a::{plus,times}^('b + 'c))"
+  by (simp add: Cart_eq)
+
+lemma sndcart_neg[simp]:"sndcart(- x) = - sndcart (x::'a::ring_1^('b + 'c))"
+  by (simp add: Cart_eq)
+
+lemma sndcart_sub[simp]:"sndcart(x - y) = sndcart (x::'a::ring_1^('b + 'c)) - sndcart y"
+  by (simp add: Cart_eq)
 
 lemma sndcart_setsum:
   fixes f:: "'d \<Rightarrow> 'a::semiring_1^_"
@@ -2533,10 +2513,10 @@
   by (induct rule: finite_induct[OF fS], simp_all add: vec_0[symmetric] del: vec_0)
 
 lemma pastecart_vec[simp]: "pastecart (vec x) (vec x) = vec x"
-  by (simp add: pastecart_eq fstcart_vec sndcart_vec fstcart_pastecart sndcart_pastecart)
+  by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart)
 
 lemma pastecart_add[simp]:"pastecart (x1::'a::{plus,times}^_) y1 + pastecart x2 y2 = pastecart (x1 + x2) (y1 + y2)"
-  by (simp add: pastecart_eq fstcart_add sndcart_add fstcart_pastecart sndcart_pastecart)
+  by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart)
 
 lemma pastecart_cmul[simp]: "pastecart (c *s (x1::'a::{plus,times}^_)) (c *s y1) = c *s pastecart x1 y1"
   by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart)
@@ -2553,109 +2533,53 @@
   shows "pastecart (setsum f S) (setsum g S) = setsum (\<lambda>i. pastecart (f i) (g i)) S"
   by (simp  add: pastecart_eq fstcart_setsum[OF fS] sndcart_setsum[OF fS] fstcart_pastecart sndcart_pastecart)
 
-lemma norm_fstcart: "norm(fstcart x) <= norm (x::real ^('n,'m) finite_sum)"
+lemma setsum_Plus:
+  "\<lbrakk>finite A; finite B\<rbrakk> \<Longrightarrow>
+    (\<Sum>x\<in>A <+> B. g x) = (\<Sum>x\<in>A. g (Inl x)) + (\<Sum>x\<in>B. g (Inr x))"
+  unfolding Plus_def
+  by (subst setsum_Un_disjoint, auto simp add: setsum_reindex)
+
+lemma setsum_UNIV_sum:
+  fixes g :: "'a::finite + 'b::finite \<Rightarrow> _"
+  shows "(\<Sum>x\<in>UNIV. g x) = (\<Sum>x\<in>UNIV. g (Inl x)) + (\<Sum>x\<in>UNIV. g (Inr x))"
+  apply (subst UNIV_Plus_UNIV [symmetric])
+  apply (rule setsum_Plus [OF finite finite])
+  done
+
+lemma norm_fstcart: "norm(fstcart x) <= norm (x::real ^('n::finite + 'm::finite))"
 proof-
-  let ?n = "dimindex (UNIV :: 'n set)"
-  let ?m = "dimindex (UNIV :: 'm set)"
-  let ?N = "{1 .. ?n}"
-  let ?M = "{1 .. ?m}"
-  let ?NM = "{1 .. dimindex (UNIV :: ('n,'m) finite_sum set)}"
-  have th_0: "1 \<le> ?n +1" by simp
   have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))"
     by (simp add: pastecart_fst_snd)
   have th1: "fstcart x \<bullet> fstcart x \<le> pastecart (fstcart x) (sndcart x) \<bullet> pastecart (fstcart x) (sndcart x)"
-    by (simp add: dot_def setsum_add_split[OF th_0, of _ ?m] pastecart_def dimindex_finite_sum Cart_lambda_beta setsum_nonneg zero_le_square del: One_nat_def)
+    by (simp add: dot_def setsum_UNIV_sum pastecart_def setsum_nonneg)
   then show ?thesis
     unfolding th0
     unfolding real_vector_norm_def real_sqrt_le_iff id_def
-    by (simp add: dot_def dimindex_finite_sum Cart_lambda_beta)
+    by (simp add: dot_def)
 qed
 
 lemma dist_fstcart: "dist(fstcart (x::real^_)) (fstcart y) <= dist x y"
   by (metis dist_def fstcart_sub[symmetric] norm_fstcart)
 
-lemma norm_sndcart: "norm(sndcart x) <= norm (x::real ^('n,'m) finite_sum)"
+lemma norm_sndcart: "norm(sndcart x) <= norm (x::real ^('n::finite + 'm::finite))"
 proof-
-  let ?n = "dimindex (UNIV :: 'n set)"
-  let ?m = "dimindex (UNIV :: 'm set)"
-  let ?N = "{1 .. ?n}"
-  let ?M = "{1 .. ?m}"
-  let ?nm = "dimindex (UNIV :: ('n,'m) finite_sum set)"
-  let ?NM = "{1 .. ?nm}"
-  have thnm[simp]: "?nm = ?n + ?m" by (simp add: dimindex_finite_sum)
-  have th_0: "1 \<le> ?n +1" by simp
   have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))"
     by (simp add: pastecart_fst_snd)
-  let ?f = "\<lambda>n. n - ?n"
-  let ?S = "{?n+1 .. ?nm}"
-  have finj:"inj_on ?f ?S"
-    using dimindex_nonzero[of "UNIV :: 'n set"] dimindex_nonzero[of "UNIV :: 'm set"]
-    apply (simp add: Ball_def atLeastAtMost_iff inj_on_def dimindex_finite_sum del: One_nat_def)
-    by arith
-  have fS: "?f ` ?S = ?M"
-    apply (rule set_ext)
-    apply (simp add: image_iff Bex_def) using dimindex_nonzero[of "UNIV :: 'n set"] dimindex_nonzero[of "UNIV :: 'm set"] by arith
   have th1: "sndcart x \<bullet> sndcart x \<le> pastecart (fstcart x) (sndcart x) \<bullet> pastecart (fstcart x) (sndcart x)"
-    by (simp add: dot_def setsum_add_split[OF th_0, of _ ?m] pastecart_def dimindex_finite_sum Cart_lambda_beta setsum_nonneg zero_le_square setsum_reindex[OF finj, unfolded fS] del: One_nat_def)
+    by (simp add: dot_def setsum_UNIV_sum pastecart_def setsum_nonneg)
   then show ?thesis
     unfolding th0
     unfolding real_vector_norm_def real_sqrt_le_iff id_def
-    by (simp add: dot_def dimindex_finite_sum Cart_lambda_beta)
+    by (simp add: dot_def)
 qed
 
 lemma dist_sndcart: "dist(sndcart (x::real^_)) (sndcart y) <= dist x y"
   by (metis dist_def sndcart_sub[symmetric] norm_sndcart)
 
-lemma dot_pastecart: "(pastecart (x1::'a::{times,comm_monoid_add}^'n) (x2::'a::{times,comm_monoid_add}^'m)) \<bullet> (pastecart y1 y2) =  x1 \<bullet> y1 + x2 \<bullet> y2"
-proof-
-  let ?n = "dimindex (UNIV :: 'n set)"
-  let ?m = "dimindex (UNIV :: 'm set)"
-  let ?N = "{1 .. ?n}"
-  let ?M = "{1 .. ?m}"
-  let ?nm = "dimindex (UNIV :: ('n,'m) finite_sum set)"
-  let ?NM = "{1 .. ?nm}"
-  have thnm: "?nm = ?n + ?m" by (simp add: dimindex_finite_sum)
-  have th_0: "1 \<le> ?n +1" by simp
-  have th_1: "\<And>i. i \<in> {?m + 1 .. ?nm} \<Longrightarrow> i - ?m \<in> ?N" apply (simp add: thnm) by arith
-  let ?f = "\<lambda>a b i. (a$i) * (b$i)"
-  let ?g = "?f (pastecart x1 x2) (pastecart y1 y2)"
-  let ?S = "{?n +1 .. ?nm}"
-  {fix i
-    assume i: "i \<in> ?N"
-    have "?g i = ?f x1 y1 i"
-      using i
-      apply (simp add: pastecart_def Cart_lambda_beta thnm) done
-  }
-  hence th2: "setsum ?g ?N = setsum (?f x1 y1) ?N"
-    apply -
-    apply (rule setsum_cong)
-    apply auto
-    done
-  {fix i
-    assume i: "i \<in> ?S"
-    have "?g i = ?f x2 y2 (i - ?n)"
-      using i
-      apply (simp add: pastecart_def Cart_lambda_beta thnm) done
-  }
-  hence th3: "setsum ?g ?S = setsum (\<lambda>i. ?f x2 y2 (i -?n)) ?S"
-    apply -
-    apply (rule setsum_cong)
-    apply auto
-    done
-  let ?r = "\<lambda>n. n - ?n"
-  have rinj: "inj_on ?r ?S" apply (simp add: inj_on_def Ball_def thnm) by arith
-  have rS: "?r ` ?S = ?M" apply (rule set_ext)
-    apply (simp add: thnm image_iff Bex_def) by arith
-  have "pastecart x1 x2 \<bullet> (pastecart y1 y2) = setsum ?g ?NM" by (simp add: dot_def)
-  also have "\<dots> = setsum ?g ?N + setsum ?g ?S"
-    by (simp add: dot_def thnm setsum_add_split[OF th_0, of _ ?m] del: One_nat_def)
-  also have "\<dots> = setsum (?f x1 y1) ?N + setsum (?f x2 y2) ?M"
-    unfolding setsum_reindex[OF rinj, unfolded rS o_def] th2 th3 ..
-  finally
-  show ?thesis by (simp add: dot_def)
-qed
-
-lemma norm_pastecart: "norm(pastecart x y) <= norm(x :: real ^ _) + norm(y)"
+lemma dot_pastecart: "(pastecart (x1::'a::{times,comm_monoid_add}^'n::finite) (x2::'a::{times,comm_monoid_add}^'m::finite)) \<bullet> (pastecart y1 y2) =  x1 \<bullet> y1 + x2 \<bullet> y2"
+  by (simp add: dot_def setsum_UNIV_sum pastecart_def)
+
+lemma norm_pastecart: "norm(pastecart x y) <= norm(x :: real ^ 'm::finite) + norm(y::real^'n::finite)"
   unfolding real_vector_norm_def dot_pastecart real_sqrt_le_iff id_def
   apply (rule power2_le_imp_le)
   apply (simp add: real_sqrt_pow2[OF add_nonneg_nonneg[OF dot_pos_le[of x] dot_pos_le[of y]]])
@@ -3419,7 +3343,7 @@
 
 (* Standard bases are a spanning set, and obviously finite.                  *)
 
-lemma span_stdbasis:"span {basis i :: 'a::ring_1^'n | i. i \<in> {1 .. dimindex(UNIV :: 'n set)}} = UNIV"
+lemma span_stdbasis:"span {basis i :: 'a::ring_1^'n::finite | i. i \<in> (UNIV :: 'n set)} = UNIV"
 apply (rule set_ext)
 apply auto
 apply (subst basis_expansion[symmetric])
@@ -3431,47 +3355,43 @@
 apply (auto simp add: Collect_def mem_def)
 done
 
-
-lemma has_size_stdbasis: "{basis i ::real ^'n | i. i \<in> {1 .. dimindex (UNIV :: 'n set)}} hassize (dimindex(UNIV :: 'n set))" (is "?S hassize ?n")
+lemma has_size_stdbasis: "{basis i ::real ^'n::finite | i. i \<in> (UNIV :: 'n set)} hassize CARD('n)" (is "?S hassize ?n")
 proof-
-  have eq: "?S = basis ` {1 .. ?n}" by blast
+  have eq: "?S = basis ` UNIV" by blast
   show ?thesis unfolding eq
     apply (rule hassize_image_inj[OF basis_inj])
     by (simp add: hassize_def)
 qed
 
-lemma finite_stdbasis: "finite {basis i ::real^'n |i. i\<in> {1 .. dimindex(UNIV:: 'n set)}}"
+lemma finite_stdbasis: "finite {basis i ::real^'n::finite |i. i\<in> (UNIV:: 'n set)}"
   using has_size_stdbasis[unfolded hassize_def]
   ..
 
-lemma card_stdbasis: "card {basis i ::real^'n |i. i\<in> {1 .. dimindex(UNIV :: 'n set)}} = dimindex(UNIV :: 'n set)"
+lemma card_stdbasis: "card {basis i ::real^'n::finite |i. i\<in> (UNIV :: 'n set)} = CARD('n)"
   using has_size_stdbasis[unfolded hassize_def]
   ..
 
 lemma independent_stdbasis_lemma:
   assumes x: "(x::'a::semiring_1 ^ 'n) \<in> span (basis ` S)"
-  and i: "i \<in> {1 .. dimindex (UNIV :: 'n set)}"
   and iS: "i \<notin> S"
   shows "(x$i) = 0"
 proof-
-  let ?n = "dimindex (UNIV :: 'n set)"
-  let ?U = "{1 .. ?n}"
+  let ?U = "UNIV :: 'n set"
   let ?B = "basis ` S"
   let ?P = "\<lambda>(x::'a^'n). \<forall>i\<in> ?U. i \<notin> S \<longrightarrow> x$i =0"
  {fix x::"'a^'n" assume xS: "x\<in> ?B"
-   from xS have "?P x" by (auto simp add: basis_component)}
+   from xS have "?P x" by auto}
  moreover
  have "subspace ?P"
-   by (auto simp add: subspace_def Collect_def mem_def zero_index vector_component)
+   by (auto simp add: subspace_def Collect_def mem_def)
  ultimately show ?thesis
-   using x span_induct[of ?B ?P x] i iS by blast
+   using x span_induct[of ?B ?P x] iS by blast
 qed
 
-lemma independent_stdbasis: "independent {basis i ::real^'n |i. i\<in> {1 .. dimindex(UNIV :: 'n set)}}"
+lemma independent_stdbasis: "independent {basis i ::real^'n::finite |i. i\<in> (UNIV :: 'n set)}"
 proof-
-  let ?n = "dimindex (UNIV :: 'n set)"
-  let ?I = "{1 .. ?n}"
-  let ?b = "basis :: nat \<Rightarrow> real ^'n"
+  let ?I = "UNIV :: 'n set"
+  let ?b = "basis :: _ \<Rightarrow> real ^'n"
   let ?B = "?b ` ?I"
   have eq: "{?b i|i. i \<in> ?I} = ?B"
     by auto
@@ -3484,8 +3404,8 @@
       apply (rule inj_on_image_set_diff[symmetric])
       apply (rule basis_inj) using k(1) by auto
     from k(2) have th0: "?b k \<in> span (?b ` (?I - {k}))" unfolding eq2 .
-    from independent_stdbasis_lemma[OF th0 k(1), simplified]
-    have False by (simp add: basis_component[OF k(1), of k])}
+    from independent_stdbasis_lemma[OF th0, of k, simplified]
+    have False by simp}
   then show ?thesis unfolding eq dependent_def ..
 qed
 
@@ -3665,19 +3585,19 @@
     done
 qed
 
-lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x\<in> {(i::nat) .. j}}"
+lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x\<in> (UNIV::'a::finite set)}"
 proof-
-  have eq: "{f x |x. x\<in> {i .. j}} = f ` {i .. j}" by auto
+  have eq: "{f x |x. x\<in> UNIV} = f ` UNIV" by auto
   show ?thesis unfolding eq
     apply (rule finite_imageI)
-    apply (rule finite_atLeastAtMost)
+    apply (rule finite)
     done
 qed
 
 
 lemma independent_bound:
-  fixes S:: "(real^'n) set"
-  shows "independent S \<Longrightarrow> finite S \<and> card S <= dimindex(UNIV :: 'n set)"
+  fixes S:: "(real^'n::finite) set"
+  shows "independent S \<Longrightarrow> finite S \<and> card S <= CARD('n)"
   apply (subst card_stdbasis[symmetric])
   apply (rule independent_span_bound)
   apply (rule finite_Atleast_Atmost_nat)
@@ -3686,23 +3606,23 @@
   apply (rule subset_UNIV)
   done
 
-lemma dependent_biggerset: "(finite (S::(real ^'n) set) ==> card S > dimindex(UNIV:: 'n set)) ==> dependent S"
+lemma dependent_biggerset: "(finite (S::(real ^'n::finite) set) ==> card S > CARD('n)) ==> dependent S"
   by (metis independent_bound not_less)
 
 (* Hence we can create a maximal independent subset.                         *)
 
 lemma maximal_independent_subset_extend:
-  assumes sv: "(S::(real^'n) set) \<subseteq> V" and iS: "independent S"
+  assumes sv: "(S::(real^'n::finite) set) \<subseteq> V" and iS: "independent S"
   shows "\<exists>B. S \<subseteq> B \<and> B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B"
   using sv iS
-proof(induct d\<equiv> "dimindex (UNIV :: 'n set) - card S" arbitrary: S rule: nat_less_induct)
+proof(induct d\<equiv> "CARD('n) - card S" arbitrary: S rule: nat_less_induct)
   fix n and S:: "(real^'n) set"
-  assume H: "\<forall>m<n. \<forall>S \<subseteq> V. independent S \<longrightarrow> m = dimindex (UNIV::'n set) - card S \<longrightarrow>
+  assume H: "\<forall>m<n. \<forall>S \<subseteq> V. independent S \<longrightarrow> m = CARD('n) - card S \<longrightarrow>
               (\<exists>B. S \<subseteq> B \<and> B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B)"
-    and sv: "S \<subseteq> V" and i: "independent S" and n: "n = dimindex (UNIV :: 'n set) - card S"
+    and sv: "S \<subseteq> V" and i: "independent S" and n: "n = CARD('n) - card S"
   let ?P = "\<lambda>B. S \<subseteq> B \<and> B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B"
   let ?ths = "\<exists>x. ?P x"
-  let ?d = "dimindex (UNIV :: 'n set)"
+  let ?d = "CARD('n)"
   {assume "V \<subseteq> span S"
     then have ?ths  using sv i by blast }
   moreover
@@ -3713,7 +3633,7 @@
     from independent_insert[of a S]  i a
     have th1: "independent (insert a S)" by auto
     have mlt: "?d - card (insert a S) < n"
-      using aS a n independent_bound[OF th1] dimindex_ge_1[of "UNIV :: 'n set"]
+      using aS a n independent_bound[OF th1]
       by auto
 
     from H[rule_format, OF mlt th0 th1 refl]
@@ -3725,14 +3645,14 @@
 qed
 
 lemma maximal_independent_subset:
-  "\<exists>(B:: (real ^'n) set). B\<subseteq> V \<and> independent B \<and> V \<subseteq> span B"
+  "\<exists>(B:: (real ^'n::finite) set). B\<subseteq> V \<and> independent B \<and> V \<subseteq> span B"
   by (metis maximal_independent_subset_extend[of "{}:: (real ^'n) set"] empty_subsetI independent_empty)
 
 (* Notion of dimension.                                                      *)
 
 definition "dim V = (SOME n. \<exists>B. B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (B hassize n))"
 
-lemma basis_exists:  "\<exists>B. (B :: (real ^'n) set) \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (B hassize dim V)"
+lemma basis_exists:  "\<exists>B. (B :: (real ^'n::finite) set) \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (B hassize dim V)"
 unfolding dim_def some_eq_ex[of "\<lambda>n. \<exists>B. B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (B hassize n)"]
 unfolding hassize_def
 using maximal_independent_subset[of V] independent_bound
@@ -3740,37 +3660,37 @@
 
 (* Consequences of independence or spanning for cardinality.                 *)
 
-lemma independent_card_le_dim: "(B::(real ^'n) set) \<subseteq> V \<Longrightarrow> independent B \<Longrightarrow> finite B \<and> card B \<le> dim V"
+lemma independent_card_le_dim: "(B::(real ^'n::finite) set) \<subseteq> V \<Longrightarrow> independent B \<Longrightarrow> finite B \<and> card B \<le> dim V"
 by (metis basis_exists[of V] independent_span_bound[where ?'a=real] hassize_def subset_trans)
 
-lemma span_card_ge_dim:  "(B::(real ^'n) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> finite B \<Longrightarrow> dim V \<le> card B"
+lemma span_card_ge_dim:  "(B::(real ^'n::finite) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> finite B \<Longrightarrow> dim V \<le> card B"
   by (metis basis_exists[of V] independent_span_bound hassize_def subset_trans)
 
 lemma basis_card_eq_dim:
-  "B \<subseteq> (V:: (real ^'n) set) \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> finite B \<and> card B = dim V"
+  "B \<subseteq> (V:: (real ^'n::finite) set) \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> finite B \<and> card B = dim V"
   by (metis order_eq_iff independent_card_le_dim span_card_ge_dim independent_mono)
 
-lemma dim_unique: "(B::(real ^'n) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> B hassize n \<Longrightarrow> dim V = n"
+lemma dim_unique: "(B::(real ^'n::finite) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> B hassize n \<Longrightarrow> dim V = n"
   by (metis basis_card_eq_dim hassize_def)
 
 (* More lemmas about dimension.                                              *)
 
-lemma dim_univ: "dim (UNIV :: (real^'n) set) = dimindex (UNIV :: 'n set)"
-  apply (rule dim_unique[of "{basis i |i. i\<in> {1 .. dimindex (UNIV :: 'n set)}}"])
+lemma dim_univ: "dim (UNIV :: (real^'n::finite) set) = CARD('n)"
+  apply (rule dim_unique[of "{basis i |i. i\<in> (UNIV :: 'n set)}"])
   by (auto simp only: span_stdbasis has_size_stdbasis independent_stdbasis)
 
 lemma dim_subset:
-  "(S:: (real ^'n) set) \<subseteq> T \<Longrightarrow> dim S \<le> dim T"
+  "(S:: (real ^'n::finite) set) \<subseteq> T \<Longrightarrow> dim S \<le> dim T"
   using basis_exists[of T] basis_exists[of S]
   by (metis independent_span_bound[where ?'a = real and ?'n = 'n] subset_eq hassize_def)
 
-lemma dim_subset_univ: "dim (S:: (real^'n) set) \<le> dimindex (UNIV :: 'n set)"
+lemma dim_subset_univ: "dim (S:: (real^'n::finite) set) \<le> CARD('n)"
   by (metis dim_subset subset_UNIV dim_univ)
 
 (* Converses to those.                                                       *)
 
 lemma card_ge_dim_independent:
-  assumes BV:"(B::(real ^'n) set) \<subseteq> V" and iB:"independent B" and dVB:"dim V \<le> card B"
+  assumes BV:"(B::(real ^'n::finite) set) \<subseteq> V" and iB:"independent B" and dVB:"dim V \<le> card B"
   shows "V \<subseteq> span B"
 proof-
   {fix a assume aV: "a \<in> V"
@@ -3784,7 +3704,7 @@
 qed
 
 lemma card_le_dim_spanning:
-  assumes BV: "(B:: (real ^'n) set) \<subseteq> V" and VB: "V \<subseteq> span B"
+  assumes BV: "(B:: (real ^'n::finite) set) \<subseteq> V" and VB: "V \<subseteq> span B"
   and fB: "finite B" and dVB: "dim V \<ge> card B"
   shows "independent B"
 proof-
@@ -3805,7 +3725,7 @@
   then show ?thesis unfolding dependent_def by blast
 qed
 
-lemma card_eq_dim: "(B:: (real ^'n) set) \<subseteq> V \<Longrightarrow> B hassize dim V \<Longrightarrow> independent B \<longleftrightarrow> V \<subseteq> span B"
+lemma card_eq_dim: "(B:: (real ^'n::finite) set) \<subseteq> V \<Longrightarrow> B hassize dim V \<Longrightarrow> independent B \<longleftrightarrow> V \<subseteq> span B"
   by (metis hassize_def order_eq_iff card_le_dim_spanning
     card_ge_dim_independent)
 
@@ -3814,13 +3734,13 @@
 (* ------------------------------------------------------------------------- *)
 
 lemma independent_bound_general:
-  "independent (S:: (real^'n) set) \<Longrightarrow> finite S \<and> card S \<le> dim S"
+  "independent (S:: (real^'n::finite) set) \<Longrightarrow> finite S \<and> card S \<le> dim S"
   by (metis independent_card_le_dim independent_bound subset_refl)
 
-lemma dependent_biggerset_general: "(finite (S:: (real^'n) set) \<Longrightarrow> card S > dim S) \<Longrightarrow> dependent S"
+lemma dependent_biggerset_general: "(finite (S:: (real^'n::finite) set) \<Longrightarrow> card S > dim S) \<Longrightarrow> dependent S"
   using independent_bound_general[of S] by (metis linorder_not_le)
 
-lemma dim_span: "dim (span (S:: (real ^'n) set)) = dim S"
+lemma dim_span: "dim (span (S:: (real ^'n::finite) set)) = dim S"
 proof-
   have th0: "dim S \<le> dim (span S)"
     by (auto simp add: subset_eq intro: dim_subset span_superset)
@@ -3833,10 +3753,10 @@
     using fB(2)  by arith
 qed
 
-lemma subset_le_dim: "(S:: (real ^'n) set) \<subseteq> span T \<Longrightarrow> dim S \<le> dim T"
+lemma subset_le_dim: "(S:: (real ^'n::finite) set) \<subseteq> span T \<Longrightarrow> dim S \<le> dim T"
   by (metis dim_span dim_subset)
 
-lemma span_eq_dim: "span (S:: (real ^'n) set) = span T ==> dim S = dim T"
+lemma span_eq_dim: "span (S:: (real ^'n::finite) set) = span T ==> dim S = dim T"
   by (metis dim_span)
 
 lemma spans_image:
@@ -3845,7 +3765,9 @@
   unfolding span_linear_image[OF lf]
   by (metis VB image_mono)
 
-lemma dim_image_le: assumes lf: "linear f" shows "dim (f ` S) \<le> dim (S:: (real ^'n) set)"
+lemma dim_image_le:
+  fixes f :: "real^'n::finite \<Rightarrow> real^'m::finite"
+  assumes lf: "linear f" shows "dim (f ` S) \<le> dim (S:: (real ^'n::finite) set)"
 proof-
   from basis_exists[of S] obtain B where
     B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "B hassize dim S" by blast
@@ -3889,14 +3811,14 @@
     (* FIXME : Move to some general theory ?*)
 definition "pairwise R S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y\<in> S. x\<noteq>y \<longrightarrow> R x y)"
 
-lemma vector_sub_project_orthogonal: "(b::'a::ordered_field^'n) \<bullet> (x - ((b \<bullet> x) / (b\<bullet>b)) *s b) = 0"
+lemma vector_sub_project_orthogonal: "(b::'a::ordered_field^'n::finite) \<bullet> (x - ((b \<bullet> x) / (b\<bullet>b)) *s b) = 0"
   apply (cases "b = 0", simp)
   apply (simp add: dot_rsub dot_rmult)
   unfolding times_divide_eq_right[symmetric]
   by (simp add: field_simps dot_eq_0)
 
 lemma basis_orthogonal:
-  fixes B :: "(real ^'n) set"
+  fixes B :: "(real ^'n::finite) set"
   assumes fB: "finite B"
   shows "\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C"
   (is " \<exists>C. ?P B C")
@@ -3972,7 +3894,7 @@
 qed
 
 lemma orthogonal_basis_exists:
-  fixes V :: "(real ^'n) set"
+  fixes V :: "(real ^'n::finite) set"
   shows "\<exists>B. independent B \<and> B \<subseteq> span V \<and> V \<subseteq> span B \<and> (B hassize dim V) \<and> pairwise orthogonal B"
 proof-
   from basis_exists[of V] obtain B where B: "B \<subseteq> V" "independent B" "V \<subseteq> span B" "B hassize dim V" by blast
@@ -4000,7 +3922,7 @@
 
 lemma span_not_univ_orthogonal:
   assumes sU: "span S \<noteq> UNIV"
-  shows "\<exists>(a:: real ^'n). a \<noteq>0 \<and> (\<forall>x \<in> span S. a \<bullet> x = 0)"
+  shows "\<exists>(a:: real ^'n::finite). a \<noteq>0 \<and> (\<forall>x \<in> span S. a \<bullet> x = 0)"
 proof-
   from sU obtain a where a: "a \<notin> span S" by blast
   from orthogonal_basis_exists obtain B where
@@ -4039,17 +3961,17 @@
 qed
 
 lemma span_not_univ_subset_hyperplane:
-  assumes SU: "span S \<noteq> (UNIV ::(real^'n) set)"
+  assumes SU: "span S \<noteq> (UNIV ::(real^'n::finite) set)"
   shows "\<exists> a. a \<noteq>0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
   using span_not_univ_orthogonal[OF SU] by auto
 
 lemma lowdim_subset_hyperplane:
-  assumes d: "dim S < dimindex (UNIV :: 'n set)"
-  shows "\<exists>(a::real ^'n). a  \<noteq> 0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
+  assumes d: "dim S < CARD('n::finite)"
+  shows "\<exists>(a::real ^'n::finite). a  \<noteq> 0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
 proof-
   {assume "span S = UNIV"
     hence "dim (span S) = dim (UNIV :: (real ^'n) set)" by simp
-    hence "dim S = dimindex (UNIV :: 'n set)" by (simp add: dim_span dim_univ)
+    hence "dim S = CARD('n)" by (simp add: dim_span dim_univ)
     with d have False by arith}
   hence th: "span S \<noteq> UNIV" by blast
   from span_not_univ_subset_hyperplane[OF th] show ?thesis .
@@ -4196,7 +4118,7 @@
 qed
 
 lemma linear_independent_extend:
-  assumes iB: "independent (B:: (real ^'n) set)"
+  assumes iB: "independent (B:: (real ^'n::finite) set)"
   shows "\<exists>g. linear g \<and> (\<forall>x\<in>B. g x = f x)"
 proof-
   from maximal_independent_subset_extend[of B UNIV] iB
@@ -4249,7 +4171,8 @@
 qed
 
 lemma subspace_isomorphism:
-  assumes s: "subspace (S:: (real ^'n) set)" and t: "subspace T"
+  assumes s: "subspace (S:: (real ^'n::finite) set)"
+  and t: "subspace (T :: (real ^ 'm::finite) set)"
   and d: "dim S = dim T"
   shows "\<exists>f. linear f \<and> f ` S = T \<and> inj_on f S"
 proof-
@@ -4324,12 +4247,12 @@
 qed
 
 lemma linear_eq_stdbasis:
-  assumes lf: "linear (f::'a::ring_1^'m \<Rightarrow> 'a^'n)" and lg: "linear g"
-  and fg: "\<forall>i \<in> {1 .. dimindex(UNIV :: 'm set)}. f (basis i) = g(basis i)"
+  assumes lf: "linear (f::'a::ring_1^'m::finite \<Rightarrow> 'a^'n::finite)" and lg: "linear g"
+  and fg: "\<forall>i. f (basis i) = g(basis i)"
   shows "f = g"
 proof-
   let ?U = "UNIV :: 'm set"
-  let ?I = "{basis i:: 'a^'m|i. i \<in> {1 .. dimindex ?U}}"
+  let ?I = "{basis i:: 'a^'m|i. i \<in> ?U}"
   {fix x assume x: "x \<in> (UNIV :: ('a^'m) set)"
     from equalityD2[OF span_stdbasis]
     have IU: " (UNIV :: ('a^'m) set) \<subseteq> span ?I" by blast
@@ -4369,12 +4292,12 @@
 qed
 
 lemma bilinear_eq_stdbasis:
-  assumes bf: "bilinear (f:: 'a::ring_1^'m \<Rightarrow> 'a^'n \<Rightarrow> 'a^'p)"
+  assumes bf: "bilinear (f:: 'a::ring_1^'m::finite \<Rightarrow> 'a^'n::finite \<Rightarrow> 'a^'p)"
   and bg: "bilinear g"
-  and fg: "\<forall>i\<in> {1 .. dimindex (UNIV :: 'm set)}. \<forall>j\<in>  {1 .. dimindex (UNIV :: 'n set)}. f (basis i) (basis j) = g (basis i) (basis j)"
+  and fg: "\<forall>i j. f (basis i) (basis j) = g (basis i) (basis j)"
   shows "f = g"
 proof-
-  from fg have th: "\<forall>x \<in> {basis i| i. i\<in> {1 .. dimindex (UNIV :: 'm set)}}. \<forall>y\<in>  {basis j |j. j \<in> {1 .. dimindex (UNIV :: 'n set)}}. f x y = g x y" by blast
+  from fg have th: "\<forall>x \<in> {basis i| i. i\<in> (UNIV :: 'm set)}. \<forall>y\<in>  {basis j |j. j \<in> (UNIV :: 'n set)}. f x y = g x y" by blast
   from bilinear_eq[OF bf bg equalityD2[OF span_stdbasis] equalityD2[OF span_stdbasis] th] show ?thesis by (blast intro: ext)
 qed
 
@@ -4389,16 +4312,14 @@
   by (metis matrix_transp_mul transp_mat transp_transp)
 
 lemma linear_injective_left_inverse:
-  assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" and fi: "inj f"
+  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite)" and fi: "inj f"
   shows "\<exists>g. linear g \<and> g o f = id"
 proof-
   from linear_independent_extend[OF independent_injective_image, OF independent_stdbasis, OF lf fi]
-  obtain h:: "real ^'m \<Rightarrow> real ^'n" where h: "linear h" " \<forall>x \<in> f ` {basis i|i. i \<in> {1 .. dimindex (UNIV::'n set)}}. h x = inv f x" by blast
+  obtain h:: "real ^'m \<Rightarrow> real ^'n" where h: "linear h" " \<forall>x \<in> f ` {basis i|i. i \<in> (UNIV::'n set)}. h x = inv f x" by blast
   from h(2)
-  have th: "\<forall>i\<in>{1..dimindex (UNIV::'n set)}. (h \<circ> f) (basis i) = id (basis i)"
+  have th: "\<forall>i. (h \<circ> f) (basis i) = id (basis i)"
     using inv_o_cancel[OF fi, unfolded stupid_ext[symmetric] id_def o_def]
-    apply auto
-    apply (erule_tac x="basis i" in allE)
     by auto
 
   from linear_eq_stdbasis[OF linear_compose[OF lf h(1)] linear_id th]
@@ -4407,14 +4328,14 @@
 qed
 
 lemma linear_surjective_right_inverse:
-  assumes lf: "linear (f:: real ^'m \<Rightarrow> real ^'n)" and sf: "surj f"
+  assumes lf: "linear (f:: real ^'m::finite \<Rightarrow> real ^'n::finite)" and sf: "surj f"
   shows "\<exists>g. linear g \<and> f o g = id"
 proof-
   from linear_independent_extend[OF independent_stdbasis]
   obtain h:: "real ^'n \<Rightarrow> real ^'m" where
-    h: "linear h" "\<forall> x\<in> {basis i| i. i\<in> {1 .. dimindex (UNIV :: 'n set)}}. h x = inv f x" by blast
+    h: "linear h" "\<forall> x\<in> {basis i| i. i\<in> (UNIV :: 'n set)}. h x = inv f x" by blast
   from h(2)
-  have th: "\<forall>i\<in>{1..dimindex (UNIV::'n set)}. (f o h) (basis i) = id (basis i)"
+  have th: "\<forall>i. (f o h) (basis i) = id (basis i)"
     using sf
     apply (auto simp add: surj_iff o_def stupid_ext[symmetric])
     apply (erule_tac x="basis i" in allE)
@@ -4426,7 +4347,7 @@
 qed
 
 lemma matrix_left_invertible_injective:
-"(\<exists>B. (B::real^'m^'n) ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x y. A *v x = A *v y \<longrightarrow> x = y)"
+"(\<exists>B. (B::real^'m^'n) ** (A::real^'n::finite^'m::finite) = mat 1) \<longleftrightarrow> (\<forall>x y. A *v x = A *v y \<longrightarrow> x = y)"
 proof-
   {fix B:: "real^'m^'n" and x y assume B: "B ** A = mat 1" and xy: "A *v x = A*v y"
     from xy have "B*v (A *v x) = B *v (A*v y)" by simp
@@ -4445,13 +4366,13 @@
 qed
 
 lemma matrix_left_invertible_ker:
-  "(\<exists>B. (B::real ^'m^'n) ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x. A *v x = 0 \<longrightarrow> x = 0)"
+  "(\<exists>B. (B::real ^'m::finite^'n::finite) ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x. A *v x = 0 \<longrightarrow> x = 0)"
   unfolding matrix_left_invertible_injective
   using linear_injective_0[OF matrix_vector_mul_linear, of A]
   by (simp add: inj_on_def)
 
 lemma matrix_right_invertible_surjective:
-"(\<exists>B. (A::real^'n^'m) ** (B::real^'m^'n) = mat 1) \<longleftrightarrow> surj (\<lambda>x. A *v x)"
+"(\<exists>B. (A::real^'n::finite^'m::finite) ** (B::real^'m^'n) = mat 1) \<longleftrightarrow> surj (\<lambda>x. A *v x)"
 proof-
   {fix B :: "real ^'m^'n"  assume AB: "A ** B = mat 1"
     {fix x :: "real ^ 'm"
@@ -4475,11 +4396,11 @@
 qed
 
 lemma matrix_left_invertible_independent_columns:
-  fixes A :: "real^'n^'m"
-  shows "(\<exists>(B::real ^'m^'n). B ** A = mat 1) \<longleftrightarrow> (\<forall>c. setsum (\<lambda>i. c i *s column i A) {1 .. dimindex(UNIV :: 'n set)} = 0 \<longrightarrow> (\<forall>i\<in> {1 .. dimindex (UNIV :: 'n set)}. c i = 0))"
+  fixes A :: "real^'n::finite^'m::finite"
+  shows "(\<exists>(B::real ^'m^'n). B ** A = mat 1) \<longleftrightarrow> (\<forall>c. setsum (\<lambda>i. c i *s column i A) (UNIV :: 'n set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
    (is "?lhs \<longleftrightarrow> ?rhs")
 proof-
-  let ?U = "{1 .. dimindex(UNIV :: 'n set)}"
+  let ?U = "UNIV :: 'n set"
   {assume k: "\<forall>x. A *v x = 0 \<longrightarrow> x = 0"
     {fix c i assume c: "setsum (\<lambda>i. c i *s column i A) ?U = 0"
       and i: "i \<in> ?U"
@@ -4487,7 +4408,7 @@
       have th0:"A *v ?x = 0"
 	using c
 	unfolding matrix_mult_vsum Cart_eq
-	by (auto simp add: vector_component zero_index setsum_component Cart_lambda_beta)
+	by auto
       from k[rule_format, OF th0] i
       have "c i = 0" by (vector Cart_eq)}
     hence ?rhs by blast}
@@ -4501,16 +4422,16 @@
 qed
 
 lemma matrix_right_invertible_independent_rows:
-  fixes A :: "real^'n^'m"
-  shows "(\<exists>(B::real^'m^'n). A ** B = mat 1) \<longleftrightarrow> (\<forall>c. setsum (\<lambda>i. c i *s row i A) {1 .. dimindex(UNIV :: 'm set)} = 0 \<longrightarrow> (\<forall>i\<in> {1 .. dimindex (UNIV :: 'm set)}. c i = 0))"
+  fixes A :: "real^'n::finite^'m::finite"
+  shows "(\<exists>(B::real^'m^'n). A ** B = mat 1) \<longleftrightarrow> (\<forall>c. setsum (\<lambda>i. c i *s row i A) (UNIV :: 'm set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
   unfolding left_invertible_transp[symmetric]
     matrix_left_invertible_independent_columns
   by (simp add: column_transp)
 
 lemma matrix_right_invertible_span_columns:
-  "(\<exists>(B::real ^'n^'m). (A::real ^'m^'n) ** B = mat 1) \<longleftrightarrow> span (columns A) = UNIV" (is "?lhs = ?rhs")
+  "(\<exists>(B::real ^'n::finite^'m::finite). (A::real ^'m^'n) ** B = mat 1) \<longleftrightarrow> span (columns A) = UNIV" (is "?lhs = ?rhs")
 proof-
-  let ?U = "{1 .. dimindex (UNIV :: 'm set)}"
+  let ?U = "UNIV :: 'm set"
   have fU: "finite ?U" by simp
   have lhseq: "?lhs \<longleftrightarrow> (\<forall>y. \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y)"
     unfolding matrix_right_invertible_surjective matrix_mult_vsum surj_def
@@ -4545,7 +4466,7 @@
 	  x: "setsum (\<lambda>i. (x$i) *s column i A) ?U = y2" by blast
 	let ?x = "(\<chi> j. if j = i then c + (x$i) else (x$j))::real^'m"
 	show "?P (c*s y1 + y2)"
-	  proof(rule exI[where x= "?x"], vector, auto simp add: i x[symmetric]Cart_lambda_beta setsum_component cond_value_iff right_distrib cond_application_beta vector_component cong del: if_weak_cong, simp only: One_nat_def[symmetric])
+	  proof(rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] cond_value_iff right_distrib cond_application_beta cong del: if_weak_cong)
 	    fix j
 	    have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j)
            else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))" using i(1)
@@ -4570,7 +4491,7 @@
 qed
 
 lemma matrix_left_invertible_span_rows:
-  "(\<exists>(B::real^'m^'n). B ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> span (rows A) = UNIV"
+  "(\<exists>(B::real^'m::finite^'n::finite). B ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> span (rows A) = UNIV"
   unfolding right_invertible_transp[symmetric]
   unfolding columns_transp[symmetric]
   unfolding matrix_right_invertible_span_columns
@@ -4579,7 +4500,7 @@
 (* An injective map real^'n->real^'n is also surjective.                       *)
 
 lemma linear_injective_imp_surjective:
-  assumes lf: "linear (f:: real ^'n \<Rightarrow> real ^'n)" and fi: "inj f"
+  assumes lf: "linear (f:: real ^'n::finite \<Rightarrow> real ^'n)" and fi: "inj f"
   shows "surj f"
 proof-
   let ?U = "UNIV :: (real ^'n) set"
@@ -4641,7 +4562,7 @@
 qed
 
 lemma linear_surjective_imp_injective:
-  assumes lf: "linear (f::real ^'n => real ^'n)" and sf: "surj f"
+  assumes lf: "linear (f::real ^'n::finite => real ^'n)" and sf: "surj f"
   shows "inj f"
 proof-
   let ?U = "UNIV :: (real ^'n) set"
@@ -4701,14 +4622,14 @@
   by (simp add: expand_fun_eq o_def id_def)
 
 lemma linear_injective_isomorphism:
-  assumes lf: "linear (f :: real^'n \<Rightarrow> real ^'n)" and fi: "inj f"
+  assumes lf: "linear (f :: real^'n::finite \<Rightarrow> real ^'n)" and fi: "inj f"
   shows "\<exists>f'. linear f' \<and> (\<forall>x. f' (f x) = x) \<and> (\<forall>x. f (f' x) = x)"
 unfolding isomorphism_expand[symmetric]
 using linear_surjective_right_inverse[OF lf linear_injective_imp_surjective[OF lf fi]] linear_injective_left_inverse[OF lf fi]
 by (metis left_right_inverse_eq)
 
 lemma linear_surjective_isomorphism:
-  assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'n)" and sf: "surj f"
+  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'n)" and sf: "surj f"
   shows "\<exists>f'. linear f' \<and> (\<forall>x. f' (f x) = x) \<and> (\<forall>x. f (f' x) = x)"
 unfolding isomorphism_expand[symmetric]
 using linear_surjective_right_inverse[OF lf sf] linear_injective_left_inverse[OF lf linear_surjective_imp_injective[OF lf sf]]
@@ -4717,7 +4638,7 @@
 (* Left and right inverses are the same for R^N->R^N.                        *)
 
 lemma linear_inverse_left:
-  assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'n)" and lf': "linear f'"
+  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'n)" and lf': "linear f'"
   shows "f o f' = id \<longleftrightarrow> f' o f = id"
 proof-
   {fix f f':: "real ^'n \<Rightarrow> real ^'n"
@@ -4735,7 +4656,7 @@
 (* Moreover, a one-sided inverse is automatically linear.                    *)
 
 lemma left_inverse_linear:
-  assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'n)" and gf: "g o f = id"
+  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'n)" and gf: "g o f = id"
   shows "linear g"
 proof-
   from gf have fi: "inj f" apply (auto simp add: inj_on_def o_def id_def stupid_ext[symmetric])
@@ -4750,7 +4671,7 @@
 qed
 
 lemma right_inverse_linear:
-  assumes lf: "linear (f:: real ^'n \<Rightarrow> real ^'n)" and gf: "f o g = id"
+  assumes lf: "linear (f:: real ^'n::finite \<Rightarrow> real ^'n)" and gf: "f o g = id"
   shows "linear g"
 proof-
   from gf have fi: "surj f" apply (auto simp add: surj_def o_def id_def stupid_ext[symmetric])
@@ -4767,7 +4688,7 @@
 (* The same result in terms of square matrices.                              *)
 
 lemma matrix_left_right_inverse:
-  fixes A A' :: "real ^'n^'n"
+  fixes A A' :: "real ^'n::finite^'n"
   shows "A ** A' = mat 1 \<longleftrightarrow> A' ** A = mat 1"
 proof-
   {fix A A' :: "real ^'n^'n" assume AA': "A ** A' = mat 1"
@@ -4796,21 +4717,20 @@
 
 lemma transp_columnvector:
  "transp(columnvector v) = rowvector v"
-  by (simp add: transp_def rowvector_def columnvector_def Cart_eq Cart_lambda_beta)
+  by (simp add: transp_def rowvector_def columnvector_def Cart_eq)
 
 lemma transp_rowvector: "transp(rowvector v) = columnvector v"
-  by (simp add: transp_def columnvector_def rowvector_def Cart_eq Cart_lambda_beta)
+  by (simp add: transp_def columnvector_def rowvector_def Cart_eq)
 
 lemma dot_rowvector_columnvector:
   "columnvector (A *v v) = A ** columnvector v"
   by (vector columnvector_def matrix_matrix_mult_def matrix_vector_mult_def)
 
-lemma dot_matrix_product: "(x::'a::semiring_1^'n) \<bullet> y = (((rowvector x ::'a^'n^1) ** (columnvector y :: 'a^1^'n))$1)$1"
-  apply (vector matrix_matrix_mult_def rowvector_def columnvector_def dot_def)
-  by (simp add: Cart_lambda_beta)
+lemma dot_matrix_product: "(x::'a::semiring_1^'n::finite) \<bullet> y = (((rowvector x ::'a^'n^1) ** (columnvector y :: 'a^1^'n))$1)$1"
+  by (vector matrix_matrix_mult_def rowvector_def columnvector_def dot_def)
 
 lemma dot_matrix_vector_mul:
-  fixes A B :: "real ^'n ^'n" and x y :: "real ^'n"
+  fixes A B :: "real ^'n::finite ^'n" and x y :: "real ^'n"
   shows "(A *v x) \<bullet> (B *v y) =
       (((rowvector x :: real^'n^1) ** ((transp A ** B) ** (columnvector y :: real ^1^'n)))$1)$1"
 unfolding dot_matrix_product transp_columnvector[symmetric]
@@ -4818,30 +4738,28 @@
 
 (* Infinity norm.                                                            *)
 
-definition "infnorm (x::real^'n) = rsup {abs(x$i) |i. i\<in> {1 .. dimindex(UNIV :: 'n set)}}"
-
-lemma numseg_dimindex_nonempty: "\<exists>i. i \<in> {1 .. dimindex (UNIV :: 'n set)}"
-  using dimindex_ge_1 by auto
+definition "infnorm (x::real^'n::finite) = rsup {abs(x$i) |i. i\<in> (UNIV :: 'n set)}"
+
+lemma numseg_dimindex_nonempty: "\<exists>i. i \<in> (UNIV :: 'n set)"
+  by auto
 
 lemma infnorm_set_image:
-  "{abs(x$i) |i. i\<in> {1 .. dimindex(UNIV :: 'n set)}} =
-  (\<lambda>i. abs(x$i)) ` {1 .. dimindex(UNIV :: 'n set)}" by blast
+  "{abs(x$i) |i. i\<in> (UNIV :: 'n set)} =
+  (\<lambda>i. abs(x$i)) ` (UNIV :: 'n set)" by blast
 
 lemma infnorm_set_lemma:
-  shows "finite {abs((x::'a::abs ^'n)$i) |i. i\<in> {1 .. dimindex(UNIV :: 'n set)}}"
-  and "{abs(x$i) |i. i\<in> {1 .. dimindex(UNIV :: 'n set)}} \<noteq> {}"
+  shows "finite {abs((x::'a::abs ^'n::finite)$i) |i. i\<in> (UNIV :: 'n set)}"
+  and "{abs(x$i) |i. i\<in> (UNIV :: 'n::finite set)} \<noteq> {}"
   unfolding infnorm_set_image
-  using dimindex_ge_1[of "UNIV :: 'n set"]
   by (auto intro: finite_imageI)
 
-lemma infnorm_pos_le: "0 \<le> infnorm x"
+lemma infnorm_pos_le: "0 \<le> infnorm (x::real^'n::finite)"
   unfolding infnorm_def
   unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma]
   unfolding infnorm_set_image
-  using dimindex_ge_1
   by auto
 
-lemma infnorm_triangle: "infnorm ((x::real^'n) + y) \<le> infnorm x + infnorm y"
+lemma infnorm_triangle: "infnorm ((x::real^'n::finite) + y) \<le> infnorm x + infnorm y"
 proof-
   have th: "\<And>x y (z::real). x - y <= z \<longleftrightarrow> x - z <= y" by arith
   have th1: "\<And>S f. f ` S = { f i| i. i \<in> S}" by blast
@@ -4857,12 +4775,12 @@
   unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma]
 
   unfolding infnorm_set_image ball_simps bex_simps
-  apply (simp add: vector_add_component)
-  apply (metis numseg_dimindex_nonempty th2)
+  apply simp
+  apply (metis th2)
   done
 qed
 
-lemma infnorm_eq_0: "infnorm x = 0 \<longleftrightarrow> (x::real ^'n) = 0"
+lemma infnorm_eq_0: "infnorm x = 0 \<longleftrightarrow> (x::real ^'n::finite) = 0"
 proof-
   have "infnorm x <= 0 \<longleftrightarrow> x = 0"
     unfolding infnorm_def
@@ -4880,9 +4798,7 @@
   apply (rule cong[of "rsup" "rsup"])
   apply blast
   apply (rule set_ext)
-  apply (auto simp add: vector_component abs_minus_cancel)
-  apply (rule_tac x="i" in exI)
-  apply (simp add: vector_component)
+  apply auto
   done
 
 lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)"
@@ -4905,16 +4821,16 @@
 lemma real_abs_infnorm: " \<bar>infnorm x\<bar> = infnorm x"
   using infnorm_pos_le[of x] by arith
 
-lemma component_le_infnorm: assumes i: "i \<in> {1 .. dimindex (UNIV :: 'n set)}"
-  shows "\<bar>x$i\<bar> \<le> infnorm (x::real^'n)"
+lemma component_le_infnorm:
+  shows "\<bar>x$i\<bar> \<le> infnorm (x::real^'n::finite)"
 proof-
-  let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
+  let ?U = "UNIV :: 'n set"
   let ?S = "{\<bar>x$i\<bar> |i. i\<in> ?U}"
   have fS: "finite ?S" unfolding image_Collect[symmetric]
     apply (rule finite_imageI) unfolding Collect_def mem_def by simp
-  have S0: "?S \<noteq> {}" using numseg_dimindex_nonempty by blast
+  have S0: "?S \<noteq> {}" by blast
   have th1: "\<And>S f. f ` S = { f i| i. i \<in> S}" by blast
-  from rsup_finite_in[OF fS S0] rsup_finite_Ub[OF fS S0] i
+  from rsup_finite_in[OF fS S0] rsup_finite_Ub[OF fS S0]
   show ?thesis unfolding infnorm_def isUb_def setle_def
     unfolding infnorm_set_image ball_simps by auto
 qed
@@ -4923,9 +4839,9 @@
   apply (subst infnorm_def)
   unfolding rsup_finite_le_iff[OF infnorm_set_lemma]
   unfolding infnorm_set_image ball_simps
-  apply (simp add: abs_mult vector_component del: One_nat_def)
-  apply (rule ballI)
-  apply (drule component_le_infnorm[of _ x])
+  apply (simp add: abs_mult)
+  apply (rule allI)
+  apply (cut_tac component_le_infnorm[of x])
   apply (rule mult_mono)
   apply auto
   done
@@ -4958,18 +4874,16 @@
   unfolding infnorm_set_image  ball_simps
   by (metis component_le_norm)
 lemma card_enum: "card {1 .. n} = n" by auto
-lemma norm_le_infnorm: "norm(x) <= sqrt(real (dimindex(UNIV ::'n set))) * infnorm(x::real ^'n)"
+lemma norm_le_infnorm: "norm(x) <= sqrt(real CARD('n)) * infnorm(x::real ^'n::finite)"
 proof-
-  let ?d = "dimindex(UNIV ::'n set)"
-  have d: "?d = card {1 .. ?d}" by auto
+  let ?d = "CARD('n)"
   have "real ?d \<ge> 0" by simp
   hence d2: "(sqrt (real ?d))^2 = real ?d"
     by (auto intro: real_sqrt_pow2)
   have th: "sqrt (real ?d) * infnorm x \<ge> 0"
-    by (simp add: dimindex_ge_1 zero_le_mult_iff real_sqrt_ge_0_iff infnorm_pos_le)
+    by (simp add: zero_le_mult_iff real_sqrt_ge_0_iff infnorm_pos_le)
   have th1: "x\<bullet>x \<le> (sqrt (real ?d) * infnorm x)^2"
     unfolding power_mult_distrib d2
-    apply (subst d)
     apply (subst power2_abs[symmetric])
     unfolding real_of_nat_def dot_def power2_eq_square[symmetric]
     apply (subst power2_abs[symmetric])
@@ -4986,7 +4900,7 @@
 
 (* Equality in Cauchy-Schwarz and triangle inequalities.                     *)
 
-lemma norm_cauchy_schwarz_eq: "(x::real ^'n) \<bullet> y = norm x * norm y \<longleftrightarrow> norm x *s y = norm y *s x" (is "?lhs \<longleftrightarrow> ?rhs")
+lemma norm_cauchy_schwarz_eq: "(x::real ^'n::finite) \<bullet> y = norm x * norm y \<longleftrightarrow> norm x *s y = norm y *s x" (is "?lhs \<longleftrightarrow> ?rhs")
 proof-
   {assume h: "x = 0"
     hence ?thesis by simp}
@@ -5012,7 +4926,9 @@
   ultimately show ?thesis by blast
 qed
 
-lemma norm_cauchy_schwarz_abs_eq: "abs(x \<bullet> y) = norm x * norm y \<longleftrightarrow>
+lemma norm_cauchy_schwarz_abs_eq:
+  fixes x y :: "real ^ 'n::finite"
+  shows "abs(x \<bullet> y) = norm x * norm y \<longleftrightarrow>
                 norm x *s y = norm y *s x \<or> norm(x) *s y = - norm y *s x" (is "?lhs \<longleftrightarrow> ?rhs")
 proof-
   have th: "\<And>(x::real) a. a \<ge> 0 \<Longrightarrow> abs x = a \<longleftrightarrow> x = a \<or> x = - a" by arith
@@ -5029,7 +4945,9 @@
   finally show ?thesis ..
 qed
 
-lemma norm_triangle_eq: "norm(x + y) = norm x + norm y \<longleftrightarrow> norm x *s y = norm y *s x"
+lemma norm_triangle_eq:
+  fixes x y :: "real ^ 'n::finite"
+  shows "norm(x + y) = norm x + norm y \<longleftrightarrow> norm x *s y = norm y *s x"
 proof-
   {assume x: "x =0 \<or> y =0"
     hence ?thesis by (cases "x=0", simp_all)}
@@ -5107,7 +5025,9 @@
   ultimately show ?thesis by blast
 qed
 
-lemma norm_cauchy_schwarz_equal: "abs(x \<bullet> y) = norm x * norm y \<longleftrightarrow> collinear {(0::real^'n),x,y}"
+lemma norm_cauchy_schwarz_equal:
+  fixes x y :: "real ^ 'n::finite"
+  shows "abs(x \<bullet> y) = norm x * norm y \<longleftrightarrow> collinear {(0::real^'n),x,y}"
 unfolding norm_cauchy_schwarz_abs_eq
 apply (cases "x=0", simp_all add: collinear_2)
 apply (cases "y=0", simp_all add: collinear_2 insert_commute)
--- a/src/HOL/Library/Eval_Witness.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Eval_Witness.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Evaluation Oracle with ML witnesses *}
 
 theory Eval_Witness
-imports Plain "~~/src/HOL/List"
+imports List Main
 begin
 
 text {* 
--- a/src/HOL/Library/Executable_Set.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Executable_Set.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/Executable_Set.thy
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 *)
 
 header {* Implementation of finite sets by lists *}
 
 theory Executable_Set
-imports Plain "~~/src/HOL/List"
+imports Main
 begin
 
 subsection {* Definitional rewrites *}
--- a/src/HOL/Library/Finite_Cartesian_Product.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Finite_Cartesian_Product.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,207 +5,83 @@
 header {* Definition of finite Cartesian product types. *}
 
 theory Finite_Cartesian_Product
-  (* imports Plain SetInterval ATP_Linkup *)
-imports Main
+imports Main (*FIXME: ATP_Linkup is only needed for metis at a few places. We could dispense of that by changing the proofs.*)
 begin
 
-  (* FIXME : ATP_Linkup is only needed for metis at a few places. We could dispense of that by changing the proofs*)
-subsection{* Dimention of sets *}
-
-definition "dimindex (S:: 'a set) = (if finite (UNIV::'a set) then card (UNIV:: 'a set) else 1)"
-
-syntax "_type_dimindex" :: "type => nat" ("(1DIM/(1'(_')))")
-translations "DIM(t)" => "CONST dimindex (CONST UNIV :: t set)"
-
-lemma dimindex_nonzero: "dimindex S \<noteq>  0"
-unfolding dimindex_def
-by (simp add: neq0_conv[symmetric] del: neq0_conv)
-
-lemma dimindex_ge_1: "dimindex S \<ge> 1"
-  using dimindex_nonzero[of S] by arith
-lemma dimindex_univ: "dimindex (S :: 'a set) = DIM('a)" by (simp add: dimindex_def)
-
 definition hassize (infixr "hassize" 12) where
   "(S hassize n) = (finite S \<and> card S = n)"
 
-lemma dimindex_unique: " (UNIV :: 'a set) hassize n ==> DIM('a) = n"
-by (simp add: dimindex_def hassize_def)
-
-
-subsection{* An indexing type parametrized by base type. *}
-
-typedef 'a finite_image = "{1 .. DIM('a)}"
-  using dimindex_ge_1 by auto
-
-lemma finite_image_image: "(UNIV :: 'a finite_image set) = Abs_finite_image ` {1 .. DIM('a)}"
-apply (auto simp add: Abs_finite_image_inverse image_def finite_image_def)
-apply (rule_tac x="Rep_finite_image x" in bexI)
-apply (simp_all add: Rep_finite_image_inverse Rep_finite_image)
-using Rep_finite_image[where ?'a = 'a]
-unfolding finite_image_def
-apply simp
-done
-
-text{* Dimension of such a type, and indexing over it. *}
-
-lemma inj_on_Abs_finite_image:
-  "inj_on (Abs_finite_image:: _ \<Rightarrow> 'a finite_image) {1 .. DIM('a)}"
-by (auto simp add: inj_on_def finite_image_def Abs_finite_image_inject[where ?'a='a])
-
-lemma has_size_finite_image: "(UNIV:: 'a finite_image set) hassize dimindex (S :: 'a set)"
-  unfolding hassize_def finite_image_image card_image[OF inj_on_Abs_finite_image[where ?'a='a]] by (auto simp add: dimindex_def)
-
 lemma hassize_image_inj: assumes f: "inj_on f S" and S: "S hassize n"
   shows "f ` S hassize n"
   using f S card_image[OF f]
     by (simp add: hassize_def inj_on_def)
 
-lemma card_finite_image: "card (UNIV:: 'a finite_image set) = dimindex(S:: 'a set)"
-using has_size_finite_image
-unfolding hassize_def by blast
-
-lemma finite_finite_image: "finite (UNIV:: 'a finite_image set)"
-using has_size_finite_image
-unfolding hassize_def by blast
-
-lemma dimindex_finite_image: "dimindex (S:: 'a finite_image set) = dimindex(T:: 'a set)"
-unfolding card_finite_image[of T, symmetric]
-by (auto simp add: dimindex_def finite_finite_image)
-
-lemma Abs_finite_image_works:
-  fixes i:: "'a finite_image"
-  shows " \<exists>!n \<in> {1 .. DIM('a)}. Abs_finite_image n = i"
-  unfolding Bex1_def Ex1_def
-  apply (rule_tac x="Rep_finite_image i" in exI)
-  using Rep_finite_image_inverse[where ?'a = 'a]
-    Rep_finite_image[where ?'a = 'a]
-  Abs_finite_image_inverse[where ?'a='a, symmetric]
-  by (auto simp add: finite_image_def)
-
-lemma Abs_finite_image_inj:
- "i \<in> {1 .. DIM('a)} \<Longrightarrow> j \<in> {1 .. DIM('a)}
-  \<Longrightarrow> (((Abs_finite_image i ::'a finite_image) = Abs_finite_image j) \<longleftrightarrow> (i = j))"
-  using Abs_finite_image_works[where ?'a = 'a]
-  by (auto simp add: atLeastAtMost_iff Bex1_def)
-
-lemma forall_Abs_finite_image:
-  "(\<forall>k:: 'a finite_image. P k) \<longleftrightarrow> (\<forall>i \<in> {1 .. DIM('a)}. P(Abs_finite_image i))"
-unfolding Ball_def atLeastAtMost_iff Ex1_def
-using Abs_finite_image_works[where ?'a = 'a, unfolded atLeastAtMost_iff Bex1_def]
-by metis
 
 subsection {* Finite Cartesian products, with indexing and lambdas. *}
 
-typedef (Cart)
+typedef (open Cart)
   ('a, 'b) "^" (infixl "^" 15)
-    = "{f:: 'b finite_image \<Rightarrow> 'a . True}" by simp
+    = "UNIV :: ('b \<Rightarrow> 'a) set"
+  morphisms Cart_nth Cart_lambda ..
 
-abbreviation dimset:: "('a ^ 'n) \<Rightarrow> nat set" where
-  "dimset a \<equiv> {1 .. DIM('n)}"
+notation Cart_nth (infixl "$" 90)
 
-definition Cart_nth :: "'a ^ 'b \<Rightarrow> nat \<Rightarrow> 'a" (infixl "$" 90) where
-  "x$i = Rep_Cart x (Abs_finite_image i)"
+notation (xsymbols) Cart_lambda (binder "\<chi>" 10)
 
 lemma stupid_ext: "(\<forall>x. f x = g x) \<longleftrightarrow> (f = g)"
   apply auto
   apply (rule ext)
   apply auto
   done
-lemma Cart_eq: "((x:: 'a ^ 'b) = y) \<longleftrightarrow> (\<forall>i\<in> dimset x. x$i = y$i)"
-  unfolding Cart_nth_def forall_Abs_finite_image[symmetric, where P = "\<lambda>i. Rep_Cart x i = Rep_Cart y i"] stupid_ext
-  using Rep_Cart_inject[of x y] ..
-
-consts Cart_lambda :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a ^ 'b"
-notation (xsymbols) Cart_lambda (binder "\<chi>" 10)
-
-defs Cart_lambda_def: "Cart_lambda g == (SOME (f:: 'a ^ 'b). \<forall>i \<in> {1 .. DIM('b)}. f$i = g i)"
 
-lemma  Cart_lambda_beta: " \<forall> i\<in> {1 .. DIM('b)}. (Cart_lambda g:: 'a ^ 'b)$i = g i"
-  unfolding Cart_lambda_def
-proof (rule someI_ex)
-  let ?p = "\<lambda>(i::nat) (k::'b finite_image). i \<in> {1 .. DIM('b)} \<and> (Abs_finite_image i = k)"
-  let ?f = "Abs_Cart (\<lambda>k. g (THE i. ?p i k)):: 'a ^ 'b"
-  let ?P = "\<lambda>f i. f$i = g i"
-  let ?Q = "\<lambda>(f::'a ^ 'b). \<forall> i \<in> {1 .. DIM('b)}. ?P f i"
-  {fix i
-    assume i: "i \<in> {1 .. DIM('b)}"
-    let ?j = "THE j. ?p j (Abs_finite_image i)"
-    from theI'[where P = "\<lambda>j. ?p (j::nat) (Abs_finite_image i :: 'b finite_image)", OF Abs_finite_image_works[of "Abs_finite_image i :: 'b finite_image", unfolded Bex1_def]]
-    have j: "?j \<in> {1 .. DIM('b)}" "(Abs_finite_image ?j :: 'b finite_image) = Abs_finite_image i" by blast+
-    from i j Abs_finite_image_inject[of i ?j, where ?'a = 'b]
-    have th: "?j = i" by (simp add: finite_image_def)
-    have "?P ?f i"
-      using th
-      by (simp add: Cart_nth_def Abs_Cart_inverse Rep_Cart_inverse Cart_def) }
-  hence th0: "?Q ?f" ..
-  with th0 show "\<exists>f. ?Q f" unfolding Ex1_def by auto
-qed
+lemma Cart_eq: "((x:: 'a ^ 'b) = y) \<longleftrightarrow> (\<forall>i. x$i = y$i)"
+  by (simp add: Cart_nth_inject [symmetric] expand_fun_eq)
 
-lemma  Cart_lambda_beta': "i\<in> {1 .. DIM('b)} \<Longrightarrow> (Cart_lambda g:: 'a ^ 'b)$i = g i"
-  using Cart_lambda_beta by blast
+lemma Cart_lambda_beta [simp]: "Cart_lambda g $ i = g i"
+  by (simp add: Cart_lambda_inverse)
 
 lemma Cart_lambda_unique:
   fixes f :: "'a ^ 'b"
-  shows "(\<forall>i\<in> {1 .. DIM('b)}. f$i = g i) \<longleftrightarrow> Cart_lambda g = f"
-  by (auto simp add: Cart_eq Cart_lambda_beta)
+  shows "(\<forall>i. f$i = g i) \<longleftrightarrow> Cart_lambda g = f"
+  by (auto simp add: Cart_eq)
 
-lemma Cart_lambda_eta: "(\<chi> i. (g$i)) = g" by (simp add: Cart_eq Cart_lambda_beta)
+lemma Cart_lambda_eta: "(\<chi> i. (g$i)) = g"
+  by (simp add: Cart_eq)
 
 text{* A non-standard sum to "paste" Cartesian products. *}
 
-typedef ('a,'b) finite_sum = "{1 .. DIM('a) + DIM('b)}"
-  apply (rule exI[where x="1"])
-  using dimindex_ge_1[of "UNIV :: 'a set"] dimindex_ge_1[of "UNIV :: 'b set"]
-  by auto
+definition pastecart :: "'a ^ 'm \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a ^ ('m + 'n)" where
+  "pastecart f g = (\<chi> i. case i of Inl a \<Rightarrow> f$a | Inr b \<Rightarrow> g$b)"
+
+definition fstcart:: "'a ^('m + 'n) \<Rightarrow> 'a ^ 'm" where
+  "fstcart f = (\<chi> i. (f$(Inl i)))"
 
-definition pastecart :: "'a ^ 'm \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a ^ ('m,'n) finite_sum" where
-  "pastecart f g = (\<chi> i. (if i <= DIM('m) then f$i else g$(i - DIM('m))))"
+definition sndcart:: "'a ^('m + 'n) \<Rightarrow> 'a ^ 'n" where
+  "sndcart f = (\<chi> i. (f$(Inr i)))"
 
-definition fstcart:: "'a ^('m, 'n) finite_sum \<Rightarrow> 'a ^ 'm" where
-  "fstcart f = (\<chi> i. (f$i))"
-
-definition sndcart:: "'a ^('m, 'n) finite_sum \<Rightarrow> 'a ^ 'n" where
-  "sndcart f = (\<chi> i. (f$(i + DIM('m))))"
+lemma nth_pastecart_Inl [simp]: "pastecart f g $ Inl a = f$a"
+  unfolding pastecart_def by simp
 
-lemma finite_sum_image: "(UNIV::('a,'b) finite_sum set) = Abs_finite_sum ` {1 .. DIM('a) + DIM('b)}"
-apply (auto  simp add: image_def)
-apply (rule_tac x="Rep_finite_sum x" in bexI)
-apply (simp add: Rep_finite_sum_inverse)
-using Rep_finite_sum[unfolded finite_sum_def, where ?'a = 'a and ?'b = 'b]
-apply (simp add: Rep_finite_sum)
-done
+lemma nth_pastecart_Inr [simp]: "pastecart f g $ Inr b = g$b"
+  unfolding pastecart_def by simp
+
+lemma nth_fstcart [simp]: "fstcart f $ i = f $ Inl i"
+  unfolding fstcart_def by simp
 
-lemma inj_on_Abs_finite_sum: "inj_on (Abs_finite_sum :: _ \<Rightarrow> ('a,'b) finite_sum) {1 .. DIM('a) + DIM('b)}"
-  using Abs_finite_sum_inject[where ?'a = 'a and ?'b = 'b]
-  by (auto simp add: inj_on_def finite_sum_def)
+lemma nth_sndtcart [simp]: "sndcart f $ i = f $ Inr i"
+  unfolding sndcart_def by simp
 
-lemma dimindex_has_size_finite_sum:
-  "(UNIV::('m,'n) finite_sum set) hassize (DIM('m) + DIM('n))"
-  by (simp add: finite_sum_image hassize_def card_image[OF inj_on_Abs_finite_sum[where ?'a = 'm and ?'b = 'n]] del: One_nat_def)
-
-lemma dimindex_finite_sum: "DIM(('m,'n) finite_sum) = DIM('m) + DIM('n)"
-  using dimindex_has_size_finite_sum[where ?'n = 'n and ?'m = 'm, unfolded hassize_def]
-  by (simp add: dimindex_def)
+lemma finite_sum_image: "(UNIV::('a + 'b) set) = range Inl \<union> range Inr"
+by (auto, case_tac x, auto)
 
 lemma fstcart_pastecart: "fstcart (pastecart (x::'a ^'m ) (y:: 'a ^ 'n)) = x"
-  by (simp add: pastecart_def fstcart_def Cart_eq Cart_lambda_beta dimindex_finite_sum)
+  by (simp add: Cart_eq)
 
 lemma sndcart_pastecart: "sndcart (pastecart (x::'a ^'m ) (y:: 'a ^ 'n)) = y"
-  by (simp add: pastecart_def sndcart_def Cart_eq Cart_lambda_beta dimindex_finite_sum)
+  by (simp add: Cart_eq)
 
 lemma pastecart_fst_snd: "pastecart (fstcart z) (sndcart z) = z"
-proof -
- {fix i
-  assume H: "i \<le> DIM('b) + DIM('c)"
-    "\<not> i \<le> DIM('b)"
-    from H have ith: "i - DIM('b) \<in> {1 .. DIM('c)}"
-      apply simp by arith
-    from H have th0: "i - DIM('b) + DIM('b) = i"
-      by simp
-  have "(\<chi> i. (z$(i + DIM('b))) :: 'a ^ 'c)$(i - DIM('b)) = z$i"
-    unfolding Cart_lambda_beta'[where g = "\<lambda> i. z$(i + DIM('b))", OF ith] th0 ..}
-thus ?thesis by (auto simp add: pastecart_def fstcart_def sndcart_def Cart_eq Cart_lambda_beta dimindex_finite_sum)
-qed
+  by (simp add: Cart_eq pastecart_def fstcart_def sndcart_def split: sum.split)
 
 lemma pastecart_eq: "(x = y) \<longleftrightarrow> (fstcart x = fstcart y) \<and> (sndcart x = sndcart y)"
   using pastecart_fst_snd[of x] pastecart_fst_snd[of y] by metis
@@ -216,53 +92,4 @@
 lemma exists_pastecart: "(\<exists>p. P p)  \<longleftrightarrow> (\<exists>x y. P (pastecart x y))"
   by (metis pastecart_fst_snd fstcart_pastecart sndcart_pastecart)
 
-text{* The finiteness lemma. *}
-
-lemma finite_cart:
- "\<forall>i \<in> {1 .. DIM('n)}. finite {x.  P i x}
-  \<Longrightarrow> finite {v::'a ^ 'n . (\<forall>i \<in> {1 .. DIM('n)}. P i (v$i))}"
-proof-
-  assume f: "\<forall>i \<in> {1 .. DIM('n)}. finite {x.  P i x}"
-  {fix n
-    assume n: "n \<le> DIM('n)"
-    have "finite {v:: 'a ^ 'n . (\<forall>i\<in> {1 .. DIM('n)}. i \<le> n \<longrightarrow> P i (v$i))
-                              \<and> (\<forall>i\<in> {1 .. DIM('n)}. n < i \<longrightarrow> v$i = (SOME x. False))}"
-      using n
-      proof(induct n)
-	case 0
-	have th0: "{v . (\<forall>i \<in> {1 .. DIM('n)}. v$i = (SOME x. False))} =
-      {(\<chi> i. (SOME x. False)::'a ^ 'n)}" by (auto simp add: Cart_lambda_beta Cart_eq)
-	with "0.prems" show ?case by auto
-      next
-	case (Suc n)
-	let ?h = "\<lambda>(x::'a,v:: 'a ^ 'n). (\<chi> i. if i = Suc n then x else v$i):: 'a ^ 'n"
-	let ?T = "{v\<Colon>'a ^ 'n.
-            (\<forall>i\<Colon>nat\<in>{1\<Colon>nat..DIM('n)}. i \<le> Suc n \<longrightarrow> P i (v$i)) \<and>
-            (\<forall>i\<Colon>nat\<in>{1\<Colon>nat..DIM('n)}.
-                Suc n < i \<longrightarrow> v$i = (SOME x\<Colon>'a. False))}"
-	let ?S = "{x::'a . P (Suc  n) x} \<times> {v:: 'a^'n. (\<forall>i \<in> {1 .. DIM('n)}. i <= n \<longrightarrow> P i (v$i)) \<and> (\<forall>i \<in> {1 .. DIM('n)}. n < i \<longrightarrow> v$i = (SOME x. False))}"
-	have th0: " ?T \<subseteq> (?h ` ?S)"
-	  using Suc.prems
-	  apply (auto simp add: image_def)
-	  apply (rule_tac x = "x$(Suc n)" in exI)
-	  apply (rule conjI)
-	  apply (rotate_tac)
-	  apply (erule ballE[where x="Suc n"])
-	  apply simp
-	  apply simp
-	  apply (rule_tac x= "\<chi> i. if i = Suc n then (SOME x:: 'a. False) else (x:: 'a ^ 'n)$i:: 'a ^ 'n" in exI)
-	  by (simp add: Cart_eq Cart_lambda_beta)
-	have th1: "finite ?S"
-	  apply (rule finite_cartesian_product)
-	  using f Suc.hyps Suc.prems by auto
-	from finite_imageI[OF th1] have th2: "finite (?h ` ?S)" .
-	from finite_subset[OF th0 th2] show ?case by blast
-      qed}
-
-  note th = this
-  from this[of "DIM('n)"] f
-  show ?thesis by auto
-qed
-
-
 end
--- a/src/HOL/Library/Formal_Power_Series.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Formal_Power_Series.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      Formal_Power_Series.thy
-    ID:
     Author:     Amine Chaieb, University of Cambridge
 *)
 
 header{* A formalization of formal power series *}
 
 theory Formal_Power_Series
-  imports Main Fact Parity
+imports Main Fact Parity
 begin
 
 subsection {* The type of formal power series*}
@@ -389,6 +388,14 @@
 
 instance fps :: (idom) idom ..
 
+instantiation fps :: (comm_ring_1) number_ring
+begin
+definition number_of_fps_def: "(number_of k::'a fps) = of_int k"
+
+instance 
+by (intro_classes, rule number_of_fps_def)
+end
+
 subsection{* Inverses of formal power series *}
 
 declare setsum_cong[fundef_cong]
--- a/src/HOL/Library/FrechetDeriv.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/FrechetDeriv.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title       : FrechetDeriv.thy
-    ID          : $Id$
     Author      : Brian Huffman
 *)
 
 header {* Frechet Derivative *}
 
 theory FrechetDeriv
-imports Lim
+imports Lim Complex_Main
 begin
 
 definition
@@ -223,8 +222,8 @@
   let ?k = "\<lambda>h. f (x + h) - f x"
   let ?Nf = "\<lambda>h. norm (?Rf h) / norm h"
   let ?Ng = "\<lambda>h. norm (?Rg (?k h)) / norm (?k h)"
-  from f interpret F!: bounded_linear "F" by (rule FDERIV_bounded_linear)
-  from g interpret G!: bounded_linear "G" by (rule FDERIV_bounded_linear)
+  from f interpret F: bounded_linear "F" by (rule FDERIV_bounded_linear)
+  from g interpret G: bounded_linear "G" by (rule FDERIV_bounded_linear)
   from F.bounded obtain kF where kF: "\<And>x. norm (F x) \<le> norm x * kF" by fast
   from G.bounded obtain kG where kG: "\<And>x. norm (G x) \<le> norm x * kG" by fast
 
--- a/src/HOL/Library/FuncSet.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/FuncSet.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/FuncSet.thy
-    ID:         $Id$
     Author:     Florian Kammueller and Lawrence C Paulson
 *)
 
 header {* Pi and Function Sets *}
 
 theory FuncSet
-imports Plain "~~/src/HOL/Hilbert_Choice"
+imports Hilbert_Choice Main
 begin
 
 definition
--- a/src/HOL/Library/Glbs.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Glbs.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,8 +1,6 @@
-(* Title:      Glbs
-   Author:     Amine Chaieb, University of Cambridge
-*)
+(* Author: Amine Chaieb, University of Cambridge *)
 
-header{*Definitions of Lower Bounds and Greatest Lower Bounds, analogous to Lubs*}
+header {* Definitions of Lower Bounds and Greatest Lower Bounds, analogous to Lubs *}
 
 theory Glbs
 imports Lubs
--- a/src/HOL/Library/Infinite_Set.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Infinite_Set.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,15 +1,13 @@
 (*  Title:      HOL/Library/Infinite_Set.thy
-    ID:         $Id$
     Author:     Stephan Merz
 *)
 
 header {* Infinite Sets and Related Concepts *}
 
 theory Infinite_Set
-imports Main "~~/src/HOL/SetInterval" "~~/src/HOL/Hilbert_Choice"
+imports Main
 begin
 
-
 subsection "Infinite Sets"
 
 text {*
--- a/src/HOL/Library/Inner_Product.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Inner_Product.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Inner Product Spaces and the Gradient Derivative *}
 
 theory Inner_Product
-imports Complex FrechetDeriv
+imports Complex_Main FrechetDeriv
 begin
 
 subsection {* Real inner product spaces *}
@@ -116,7 +116,7 @@
 
 end
 
-interpretation inner!:
+interpretation inner:
   bounded_bilinear "inner::'a::real_inner \<Rightarrow> 'a \<Rightarrow> real"
 proof
   fix x y z :: 'a and r :: real
@@ -135,11 +135,11 @@
   qed
 qed
 
-interpretation inner_left!:
+interpretation inner_left:
   bounded_linear "\<lambda>x::'a::real_inner. inner x y"
   by (rule inner.bounded_linear_left)
 
-interpretation inner_right!:
+interpretation inner_right:
   bounded_linear "\<lambda>y::'a::real_inner. inner x y"
   by (rule inner.bounded_linear_right)
 
--- a/src/HOL/Library/LaTeXsugar.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/LaTeXsugar.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 
 (*<*)
 theory LaTeXsugar
-imports Plain "~~/src/HOL/List"
+imports Main
 begin
 
 (* LOGIC *)
--- a/src/HOL/Library/ListVector.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/ListVector.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,11 +1,9 @@
-(*  ID:         $Id$
-    Author:     Tobias Nipkow, 2007
-*)
+(*  Author: Tobias Nipkow, 2007 *)
 
-header "Lists as vectors"
+header {* Lists as vectors *}
 
 theory ListVector
-imports Plain "~~/src/HOL/List"
+imports List Main
 begin
 
 text{* \noindent
--- a/src/HOL/Library/List_Prefix.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/List_Prefix.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/List_Prefix.thy
-    ID:         $Id$
     Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
 *)
 
 header {* List prefixes and postfixes *}
 
 theory List_Prefix
-imports Plain "~~/src/HOL/List"
+imports List Main
 begin
 
 subsection {* Prefix order on lists *}
--- a/src/HOL/Library/List_lexord.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/List_lexord.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/List_lexord.thy
-    ID:         $Id$
     Author:     Norbert Voelker
 *)
 
 header {* Lexicographic order on lists *}
 
 theory List_lexord
-imports Plain "~~/src/HOL/List"
+imports List Main
 begin
 
 instantiation list :: (ord) ord
--- a/src/HOL/Library/Mapping.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Mapping.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* An abstract view on maps for code generation. *}
 
 theory Mapping
-imports Map
+imports Map Main
 begin
 
 subsection {* Type definition and primitive operations *}
--- a/src/HOL/Library/Multiset.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Multiset.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Multisets *}
 
 theory Multiset
-imports Plain "~~/src/HOL/List"
+imports List Main
 begin
 
 subsection {* The type of multisets *}
@@ -1077,15 +1077,15 @@
 apply simp
 done
 
-interpretation mset_order!: order "op \<le>#" "op <#"
+interpretation mset_order: order "op \<le>#" "op <#"
 proof qed (auto intro: order.intro mset_le_refl mset_le_antisym
   mset_le_trans simp: mset_less_def)
 
-interpretation mset_order_cancel_semigroup!:
+interpretation mset_order_cancel_semigroup:
   pordered_cancel_ab_semigroup_add "op +" "op \<le>#" "op <#"
 proof qed (erule mset_le_mono_add [OF mset_le_refl])
 
-interpretation mset_order_semigroup_cancel!:
+interpretation mset_order_semigroup_cancel:
   pordered_ab_semigroup_add_imp_le "op +" "op \<le>#" "op <#"
 proof qed simp
 
@@ -1433,7 +1433,7 @@
 definition [code del]:
  "image_mset f = fold_mset (op + o single o f) {#}"
 
-interpretation image_left_comm!: left_commutative "op + o single o f"
+interpretation image_left_comm: left_commutative "op + o single o f"
   proof qed (simp add:union_ac)
 
 lemma image_mset_empty [simp]: "image_mset f {#} = {#}"
@@ -1623,8 +1623,8 @@
     msetT=msetT, mk_mset=mk_mset, mset_regroup_conv=regroup_munion_conv,
     mset_member_tac=mset_member_tac, mset_nonempty_tac=mset_nonempty_tac,
     mset_pwleq_tac=unfold_pwleq_tac, set_of_simps=set_of_simps,
-    smsI'=@{thm ms_strictI}, wmsI2''=@{thm ms_weakI2}, wmsI1=@{thm ms_weakI1},
-    reduction_pair=@{thm ms_reduction_pair}
+    smsI'= @{thm ms_strictI}, wmsI2''= @{thm ms_weakI2}, wmsI1= @{thm ms_weakI1},
+    reduction_pair= @{thm ms_reduction_pair}
   })
 end
 *}
--- a/src/HOL/Library/Nat_Infinity.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Nat_Infinity.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Natural numbers with infinity *}
 
 theory Nat_Infinity
-imports Plain "~~/src/HOL/Presburger"
+imports Main
 begin
 
 subsection {* Type definition *}
--- a/src/HOL/Library/Nat_Int_Bij.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Nat_Int_Bij.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Nat_Int_Bij.thy
-    ID:         $Id$
     Author:     Stefan Richter, Tobias Nipkow
 *)
 
 header{* Bijections $\mathbb{N}\to\mathbb{N}^2$ and $\mathbb{N}\to\mathbb{Z}$*}
 
 theory Nat_Int_Bij
-imports Hilbert_Choice Presburger
+imports Main
 begin
 
 subsection{*  A bijection between @{text "\<nat>"} and @{text "\<nat>\<twosuperior>"} *}
--- a/src/HOL/Library/Nested_Environment.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Nested_Environment.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/Nested_Environment.thy
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 *)
 
 header {* Nested environments *}
 
 theory Nested_Environment
-imports Plain "~~/src/HOL/List" "~~/src/HOL/Code_Eval"
+imports Main
 begin
 
 text {*
--- a/src/HOL/Library/Numeral_Type.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Numeral_Type.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Numeral Syntax for Types *}
 
 theory Numeral_Type
-imports Plain "~~/src/HOL/Presburger"
+imports Main
 begin
 
 subsection {* Preliminary lemmas *}
@@ -313,7 +313,7 @@
 
 end
 
-interpretation bit0!:
+interpretation bit0:
   mod_type "int CARD('a::finite bit0)"
            "Rep_bit0 :: 'a::finite bit0 \<Rightarrow> int"
            "Abs_bit0 :: int \<Rightarrow> 'a::finite bit0"
@@ -329,7 +329,7 @@
 apply (rule power_bit0_def [unfolded Abs_bit0'_def])
 done
 
-interpretation bit1!:
+interpretation bit1:
   mod_type "int CARD('a::finite bit1)"
            "Rep_bit1 :: 'a::finite bit1 \<Rightarrow> int"
            "Abs_bit1 :: int \<Rightarrow> 'a::finite bit1"
@@ -363,13 +363,13 @@
 
 end
 
-interpretation bit0!:
+interpretation bit0:
   mod_ring "int CARD('a::finite bit0)"
            "Rep_bit0 :: 'a::finite bit0 \<Rightarrow> int"
            "Abs_bit0 :: int \<Rightarrow> 'a::finite bit0"
   ..
 
-interpretation bit1!:
+interpretation bit1:
   mod_ring "int CARD('a::finite bit1)"
            "Rep_bit1 :: 'a::finite bit1 \<Rightarrow> int"
            "Abs_bit1 :: int \<Rightarrow> 'a::finite bit1"
--- a/src/HOL/Library/Option_ord.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Option_ord.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,15 +1,14 @@
 (*  Title:      HOL/Library/Option_ord.thy
-    ID:         $Id$
     Author:     Florian Haftmann, TU Muenchen
 *)
 
 header {* Canonical order on option type *}
 
 theory Option_ord
-imports Plain
+imports Option Main
 begin
 
-instantiation option :: (order) order
+instantiation option :: (preorder) preorder
 begin
 
 definition less_eq_option where
@@ -48,12 +47,63 @@
 lemma less_option_Some [simp, code]: "Some x < Some y \<longleftrightarrow> x < y"
   by (simp add: less_option_def)
 
-instance by default
-  (auto simp add: less_eq_option_def less_option_def split: option.splits)
+instance proof
+qed (auto simp add: less_eq_option_def less_option_def less_le_not_le elim: order_trans split: option.splits)
 
 end 
 
-instance option :: (linorder) linorder
-  by default (auto simp add: less_eq_option_def less_option_def split: option.splits)
+instance option :: (order) order proof
+qed (auto simp add: less_eq_option_def less_option_def split: option.splits)
+
+instance option :: (linorder) linorder proof
+qed (auto simp add: less_eq_option_def less_option_def split: option.splits)
+
+instantiation option :: (preorder) bot
+begin
+
+definition "bot = None"
+
+instance proof
+qed (simp add: bot_option_def)
+
+end
+
+instantiation option :: (top) top
+begin
+
+definition "top = Some top"
+
+instance proof
+qed (simp add: top_option_def less_eq_option_def split: option.split)
 
 end
+
+instance option :: (wellorder) wellorder proof
+  fix P :: "'a option \<Rightarrow> bool" and z :: "'a option"
+  assume H: "\<And>x. (\<And>y. y < x \<Longrightarrow> P y) \<Longrightarrow> P x"
+  have "P None" by (rule H) simp
+  then have P_Some [case_names Some]:
+    "\<And>z. (\<And>x. z = Some x \<Longrightarrow> (P o Some) x) \<Longrightarrow> P z"
+  proof -
+    fix z
+    assume "\<And>x. z = Some x \<Longrightarrow> (P o Some) x"
+    with `P None` show "P z" by (cases z) simp_all
+  qed
+  show "P z" proof (cases z rule: P_Some)
+    case (Some w)
+    show "(P o Some) w" proof (induct rule: less_induct)
+      case (less x)
+      have "P (Some x)" proof (rule H)
+        fix y :: "'a option"
+        assume "y < Some x"
+        show "P y" proof (cases y rule: P_Some)
+          case (Some v) with `y < Some x` have "v < x" by simp
+          with less show "(P o Some) v" .
+        qed
+      qed
+      then show ?case by simp
+    qed
+  qed
+qed
+
+end
--- a/src/HOL/Library/OptionalSugar.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/OptionalSugar.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -4,7 +4,7 @@
 *)
 (*<*)
 theory OptionalSugar
-imports LaTeXsugar Complex_Main
+imports Complex_Main LaTeXsugar
 begin
 
 (* hiding set *)
--- a/src/HOL/Library/Order_Relation.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Order_Relation.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,6 +1,4 @@
-(*  ID          : $Id$
-    Author      : Tobias Nipkow
-*)
+(* Author: Tobias Nipkow *)
 
 header {* Orders as Relations *}
 
--- a/src/HOL/Library/Permutation.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Permutation.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Permutations *}
 
 theory Permutation
-imports Plain Multiset
+imports Main Multiset
 begin
 
 inductive
@@ -188,7 +188,11 @@
    apply (subgoal_tac "set (a#list) = set (ysa@a#zs) & distinct (a#list) & distinct (ysa@a#zs)")
     apply (fastsimp simp add: insert_ident)
    apply (metis distinct_remdups set_remdups)
-  apply (metis le_less_trans Suc_length_conv length_remdups_leq less_Suc_eq nat_less_le)
+   apply (subgoal_tac "length (remdups xs) < Suc (length xs)")
+   apply simp
+   apply (subgoal_tac "length (remdups xs) \<le> length xs")
+   apply simp
+   apply (rule length_remdups_leq)
   done
 
 lemma perm_remdups_iff_eq_set: "remdups x <~~> remdups y = (set x = set y)"
--- a/src/HOL/Library/Permutations.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Permutations.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Permutations, both general and specifically on finite sets.*}
 
 theory Permutations
-imports Main Finite_Cartesian_Product Parity Fact
+imports Finite_Cartesian_Product Parity Fact Main
 begin
 
   (* Why should I import Main just to solve the Typerep problem! *)
--- a/src/HOL/Library/Pocklington.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Pocklington.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,13 +1,11 @@
 (*  Title:      HOL/Library/Pocklington.thy
-    ID:         $Id$
     Author:     Amine Chaieb
 *)
 
 header {* Pocklington's Theorem for Primes *}
 
-
 theory Pocklington
-imports Plain "~~/src/HOL/List" "~~/src/HOL/Primes"
+imports Main Primes
 begin
 
 definition modeq:: "nat => nat => nat => bool"    ("(1[_ = _] '(mod _'))")
--- a/src/HOL/Library/Polynomial.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Polynomial.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -6,7 +6,7 @@
 header {* Univariate Polynomials *}
 
 theory Polynomial
-imports Plain SetInterval Main
+imports Main
 begin
 
 subsection {* Definition of type @{text poly} *}
--- a/src/HOL/Library/Primes.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Primes.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -6,7 +6,7 @@
 header {* Primality on nat *}
 
 theory Primes
-imports Plain "~~/src/HOL/ATP_Linkup" "~~/src/HOL/GCD" "~~/src/HOL/Parity"
+imports Complex_Main
 begin
 
 definition
--- a/src/HOL/Library/Product_Vector.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Product_Vector.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -116,14 +116,14 @@
 
 subsection {* Pair operations are linear and continuous *}
 
-interpretation fst!: bounded_linear fst
+interpretation fst: bounded_linear fst
   apply (unfold_locales)
   apply (rule fst_add)
   apply (rule fst_scaleR)
   apply (rule_tac x="1" in exI, simp add: norm_Pair)
   done
 
-interpretation snd!: bounded_linear snd
+interpretation snd: bounded_linear snd
   apply (unfold_locales)
   apply (rule snd_add)
   apply (rule snd_scaleR)
--- a/src/HOL/Library/Product_ord.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Product_ord.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/Product_ord.thy
-    ID:         $Id$
     Author:     Norbert Voelker
 *)
 
 header {* Order on product types *}
 
 theory Product_ord
-imports Plain
+imports Main
 begin
 
 instantiation "*" :: (ord, ord) ord
--- a/src/HOL/Library/Quicksort.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Quicksort.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
-(*  ID:         $Id$
-    Author:     Tobias Nipkow
+(*  Author:     Tobias Nipkow
     Copyright   1994 TU Muenchen
 *)
 
 header{*Quicksort*}
 
 theory Quicksort
-imports Plain Multiset
+imports Main Multiset
 begin
 
 context linorder
--- a/src/HOL/Library/Quotient.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Quotient.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/Quotient.thy
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 *)
 
 header {* Quotient types *}
 
 theory Quotient
-imports Plain "~~/src/HOL/Hilbert_Choice"
+imports Main
 begin
 
 text {*
--- a/src/HOL/Library/RBT.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/RBT.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      RBT.thy
-    ID:         $Id$
     Author:     Markus Reiter, TU Muenchen
     Author:     Alexander Krauss, TU Muenchen
 *)
@@ -8,7 +7,7 @@
 
 (*<*)
 theory RBT
-imports Plain AssocList
+imports Main AssocList
 begin
 
 datatype color = R | B
--- a/src/HOL/Library/Ramsey.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Ramsey.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Library/Ramsey.thy
-    ID:         $Id$
     Author:     Tom Ridge. Converted to structured Isar by L C Paulson
 *)
 
 header "Ramsey's Theorem"
 
 theory Ramsey
-imports Plain "~~/src/HOL/Presburger" Infinite_Set
+imports Main Infinite_Set
 begin
 
 subsection {* Preliminaries *}
--- a/src/HOL/Library/SetsAndFunctions.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/SetsAndFunctions.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Operations on sets and functions *}
 
 theory SetsAndFunctions
-imports Plain
+imports Main
 begin
 
 text {*
@@ -107,26 +107,26 @@
   apply simp
   done
 
-interpretation set_semigroup_add!: semigroup_add "op \<oplus> :: ('a::semigroup_add) set => 'a set => 'a set"
+interpretation set_semigroup_add: semigroup_add "op \<oplus> :: ('a::semigroup_add) set => 'a set => 'a set"
   apply default
   apply (unfold set_plus_def)
   apply (force simp add: add_assoc)
   done
 
-interpretation set_semigroup_mult!: semigroup_mult "op \<otimes> :: ('a::semigroup_mult) set => 'a set => 'a set"
+interpretation set_semigroup_mult: semigroup_mult "op \<otimes> :: ('a::semigroup_mult) set => 'a set => 'a set"
   apply default
   apply (unfold set_times_def)
   apply (force simp add: mult_assoc)
   done
 
-interpretation set_comm_monoid_add!: comm_monoid_add "{0}" "op \<oplus> :: ('a::comm_monoid_add) set => 'a set => 'a set"
+interpretation set_comm_monoid_add: comm_monoid_add "{0}" "op \<oplus> :: ('a::comm_monoid_add) set => 'a set => 'a set"
   apply default
    apply (unfold set_plus_def)
    apply (force simp add: add_ac)
   apply force
   done
 
-interpretation set_comm_monoid_mult!: comm_monoid_mult "{1}" "op \<otimes> :: ('a::comm_monoid_mult) set => 'a set => 'a set"
+interpretation set_comm_monoid_mult: comm_monoid_mult "{1}" "op \<otimes> :: ('a::comm_monoid_mult) set => 'a set => 'a set"
   apply default
    apply (unfold set_times_def)
    apply (force simp add: mult_ac)
--- a/src/HOL/Library/State_Monad.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/State_Monad.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Combinator syntax for generic, open state monads (single threaded monads) *}
 
 theory State_Monad
-imports Plain "~~/src/HOL/List"
+imports Main
 begin
 
 subsection {* Motivation *}
--- a/src/HOL/Library/Sublist_Order.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Sublist_Order.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,13 +1,12 @@
 (*  Title:      HOL/Library/Sublist_Order.thy
-    ID:         $Id$
     Authors:    Peter Lammich, Uni Muenster <peter.lammich@uni-muenster.de>
-                Florian Haftmann, TU München
+                Florian Haftmann, TU Muenchen
 *)
 
 header {* Sublist Ordering *}
 
 theory Sublist_Order
-imports Plain "~~/src/HOL/List"
+imports Main
 begin
 
 text {*
--- a/src/HOL/Library/Topology_Euclidean_Space.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Topology_Euclidean_Space.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -6,10 +6,9 @@
 header {* Elementary topology in Euclidean space. *}
 
 theory Topology_Euclidean_Space
-  imports SEQ Euclidean_Space
+imports SEQ Euclidean_Space
 begin
 
-
 declare fstcart_pastecart[simp] sndcart_pastecart[simp]
 
 subsection{* General notion of a topology *}
@@ -474,7 +473,7 @@
   have th0: "\<And>d x y z. (d x z :: real) <= d x y + d y z \<Longrightarrow> d y z = d z y
                ==> ~(d x y * 2 < d x z \<and> d z y * 2 < d x z)" by arith
   have "?P ?U ?V" using dist_pos_lt[OF xy] th0[of dist,OF dist_triangle dist_sym]
-    by (auto simp add: dist_refl expand_set_eq Arith_Tools.less_divide_eq_number_of1)
+    by (auto simp add: dist_refl expand_set_eq less_divide_eq_number_of1)
   then show ?thesis by blast
 qed
 
@@ -488,7 +487,7 @@
 
 subsection{* Limit points *}
 
-definition islimpt:: "real ^'n \<Rightarrow> (real^'n) set \<Rightarrow> bool" (infixr "islimpt" 60) where
+definition islimpt:: "real ^'n::finite \<Rightarrow> (real^'n) set \<Rightarrow> bool" (infixr "islimpt" 60) where
   islimpt_def: "x islimpt S \<longleftrightarrow> (\<forall>T. x\<in>T \<longrightarrow> open T \<longrightarrow> (\<exists>y\<in>S. y\<in>T \<and> y\<noteq>x))"
 
   (* FIXME: Sure this form is OK????*)
@@ -510,7 +509,7 @@
   using approachable_lt_le[where f="\<lambda>x'. dist x' x" and P="\<lambda>x'. \<not> (x'\<in>S \<and> x'\<noteq>x)"]
   by metis
 
-lemma islimpt_UNIV[simp, intro]: "(x:: real ^'n) islimpt UNIV"
+lemma islimpt_UNIV[simp, intro]: "(x:: real ^'n::finite) islimpt UNIV"
 proof-
   {
     fix e::real assume ep: "e>0"
@@ -532,20 +531,20 @@
 lemma islimpt_EMPTY[simp]: "\<not> x islimpt {}"
   unfolding islimpt_approachable apply auto by ferrack
 
-lemma closed_positive_orthant: "closed {x::real^'n. \<forall>i\<in>{1.. dimindex(UNIV:: 'n set)}. 0 \<le>x$i}"
+lemma closed_positive_orthant: "closed {x::real^'n::finite. \<forall>i. 0 \<le>x$i}"
 proof-
-  let ?U = "{1 .. dimindex(UNIV :: 'n set)}"
-  let ?O = "{x::real^'n. \<forall>i\<in>?U. x$i\<ge>0}"
-  {fix x:: "real^'n" and i::nat assume H: "\<forall>e>0. \<exists>x'\<in>?O. x' \<noteq> x \<and> dist x' x < e" and i: "i \<in> ?U"
+  let ?U = "UNIV :: 'n set"
+  let ?O = "{x::real^'n. \<forall>i. x$i\<ge>0}"
+  {fix x:: "real^'n" and i::'n assume H: "\<forall>e>0. \<exists>x'\<in>?O. x' \<noteq> x \<and> dist x' x < e"
     and xi: "x$i < 0"
     from xi have th0: "-x$i > 0" by arith
     from H[rule_format, OF th0] obtain x' where x': "x' \<in>?O" "x' \<noteq> x" "dist x' x < -x $ i" by blast
       have th:" \<And>b a (x::real). abs x <= b \<Longrightarrow> b <= a ==> ~(a + x < 0)" by arith
       have th': "\<And>x (y::real). x < 0 \<Longrightarrow> 0 <= y ==> abs x <= abs (y - x)" by arith
-      have th1: "\<bar>x$i\<bar> \<le> \<bar>(x' - x)$i\<bar>" using i x'(1) xi
+      have th1: "\<bar>x$i\<bar> \<le> \<bar>(x' - x)$i\<bar>" using x'(1) xi
 	apply (simp only: vector_component)
 	by (rule th') auto
-      have th2: "\<bar>dist x x'\<bar> \<ge> \<bar>(x' - x)$i\<bar>" using  component_le_norm[OF i, of "x'-x"]
+      have th2: "\<bar>dist x x'\<bar> \<ge> \<bar>(x' - x)$i\<bar>" using  component_le_norm[of "x'-x" i]
 	apply (simp add: dist_def) by norm
       from th[OF th1 th2] x'(3) have False by (simp add: dist_sym dist_pos_le) }
   then show ?thesis unfolding closed_limpt islimpt_approachable
@@ -662,7 +661,7 @@
 	have "?k/2 \<ge> 0" using kp by simp
 	then obtain w where w: "dist y w = ?k/ 2" by (metis vector_choose_dist)
 	from iT[unfolded expand_set_eq mem_interior]
-	have "\<not> ball w (?k/4) \<subseteq> T" using kp by (auto simp add: Arith_Tools.less_divide_eq_number_of1)
+	have "\<not> ball w (?k/4) \<subseteq> T" using kp by (auto simp add: less_divide_eq_number_of1)
 	then obtain z where z: "dist w z < ?k/4" "z \<notin> T" by (auto simp add: subset_eq)
 	have "z \<notin> T \<and> z\<noteq> y \<and> dist z y < d \<and> dist x z < e" using z apply simp
 	  using w e(1) d apply (auto simp only: dist_sym)
@@ -965,7 +964,7 @@
 definition within :: "'a net \<Rightarrow> 'a set \<Rightarrow> 'a net" (infixr "within" 70) where
   within_def: "net within S = mknet (\<lambda>x y. netord net x y \<and> x \<in> S)"
 
-definition indirection :: "real ^'n \<Rightarrow> real ^'n \<Rightarrow> (real ^'n) net" (infixr "indirection" 70) where
+definition indirection :: "real ^'n::finite \<Rightarrow> real ^'n \<Rightarrow> (real ^'n) net" (infixr "indirection" 70) where
   indirection_def: "a indirection v = (at a) within {b. \<exists>c\<ge>0. b - a = c*s v}"
 
 text{* Prove That They are all nets. *}
@@ -1019,7 +1018,7 @@
   (\<forall>(a::'a) b. a = b) \<or> (\<exists>(a::'a) b. a \<noteq> b \<and> (\<forall>x. ~(netord (net) x a) \<and> ~(netord(net) x b)))"
 
 
-lemma trivial_limit_within: "trivial_limit (at (a::real^'n) within S) \<longleftrightarrow> ~(a islimpt S)"
+lemma trivial_limit_within: "trivial_limit (at (a::real^'n::finite) within S) \<longleftrightarrow> ~(a islimpt S)"
 proof-
   {assume "\<forall>(a::real^'n) b. a = b" hence "\<not> a islimpt S"
       apply (simp add: islimpt_approachable_le)
@@ -1104,7 +1103,7 @@
 apply (metis dlo_simps(7) dlo_simps(9) le_maxI2 min_max.le_iff_sup min_max.sup_absorb1 order_antisym_conv) done
 
 (* FIXME Declare this with P::'a::some_type \<Rightarrow> bool *)
-lemma eventually_at_infinity: "eventually (P::(real^'n \<Rightarrow> bool)) at_infinity \<longleftrightarrow> (\<exists>b. \<forall>x. norm x >= b \<longrightarrow> P x)" (is "?lhs = ?rhs")
+lemma eventually_at_infinity: "eventually (P::(real^'n::finite \<Rightarrow> bool)) at_infinity \<longleftrightarrow> (\<exists>b. \<forall>x. norm x >= b \<longrightarrow> P x)" (is "?lhs = ?rhs")
 proof
   assume "?lhs" thus "?rhs"
     unfolding eventually_def at_infinity
@@ -1145,7 +1144,7 @@
 
 subsection{* Limits, defined as vacuously true when the limit is trivial. *}
 
-definition tendsto:: "('a \<Rightarrow> real ^'n) \<Rightarrow> real ^'n \<Rightarrow> 'a net \<Rightarrow> bool" (infixr "--->" 55) where
+definition tendsto:: "('a \<Rightarrow> real ^'n::finite) \<Rightarrow> real ^'n \<Rightarrow> 'a net \<Rightarrow> bool" (infixr "--->" 55) where
   tendsto_def: "(f ---> l) net  \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) net)"
 
 lemma tendstoD: "(f ---> l) net \<Longrightarrow> e>0 \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) net"
@@ -1177,7 +1176,7 @@
   by (auto simp add: tendsto_def eventually_at)
 
 lemma Lim_at_infinity:
-  "(f ---> l) at_infinity \<longleftrightarrow> (\<forall>e>0. \<exists>b. \<forall>x::real^'n. norm x >= b \<longrightarrow> dist (f x) l < e)"
+  "(f ---> l) at_infinity \<longleftrightarrow> (\<forall>e>0. \<exists>b. \<forall>x::real^'n::finite. norm x >= b \<longrightarrow> dist (f x) l < e)"
   by (auto simp add: tendsto_def eventually_at_infinity)
 
 lemma Lim_sequentially:
@@ -1210,7 +1209,7 @@
 qed
 
 lemma Lim_Un_univ:
- "(f ---> l) (at x within S) \<Longrightarrow> (f ---> l) (at x within T) \<Longrightarrow>  S \<union> T = (UNIV::(real^'n) set)
+ "(f ---> l) (at x within S) \<Longrightarrow> (f ---> l) (at x within T) \<Longrightarrow>  S \<union> T = (UNIV::(real^'n::finite) set)
         ==> (f ---> l) (at x)"
   by (metis Lim_Un within_UNIV)
 
@@ -1275,7 +1274,7 @@
 
 text{* Basic arithmetical combining theorems for limits. *}
 
-lemma Lim_linear: fixes f :: "('a \<Rightarrow> real^'n)" and h :: "(real^'n \<Rightarrow> real^'m)"
+lemma Lim_linear: fixes f :: "('a \<Rightarrow> real^'n::finite)" and h :: "(real^'n \<Rightarrow> real^'m::finite)"
   assumes "(f ---> l) net" "linear h"
   shows "((\<lambda>x. h (f x)) ---> h l) net"
 proof (cases "trivial_limit net")
@@ -1315,7 +1314,7 @@
   apply (subst minus_diff_eq[symmetric])
   unfolding norm_minus_cancel by simp
 
-lemma Lim_add: fixes f :: "'a \<Rightarrow> real^'n" shows
+lemma Lim_add: fixes f :: "'a \<Rightarrow> real^'n::finite" shows
  "(f ---> l) net \<Longrightarrow> (g ---> m) net \<Longrightarrow> ((\<lambda>x. f(x) + g(x)) ---> l + m) net"
 proof-
   assume as:"(f ---> l) net" "(g ---> m) net"
@@ -1323,7 +1322,7 @@
     assume "e>0"
     hence *:"eventually (\<lambda>x. dist (f x) l < e/2) net"
             "eventually (\<lambda>x. dist (g x) m < e/2) net" using as
-      by (auto intro: tendstoD simp del: Arith_Tools.less_divide_eq_number_of1)
+      by (auto intro: tendstoD simp del: less_divide_eq_number_of1)
     hence "eventually (\<lambda>x. dist (f x + g x) (l + m) < e) net"
     proof(cases "trivial_limit net")
       case True
@@ -1368,14 +1367,14 @@
     using assms `e>0` unfolding tendsto_def by auto
 qed
 
-lemma Lim_component: "(f ---> l) net \<Longrightarrow> i \<in> {1 .. dimindex(UNIV:: 'n set)}
-                      ==> ((\<lambda>a. vec1((f a :: real ^'n)$i)) ---> vec1(l$i)) net"
-  apply (simp add: Lim dist_def vec1_sub[symmetric] norm_vec1  vector_minus_component[symmetric] del: One_nat_def)
-  apply auto
+lemma Lim_component: "(f ---> l) net
+                      ==> ((\<lambda>a. vec1((f a :: real ^'n::finite)$i)) ---> vec1(l$i)) net"
+  apply (simp add: Lim dist_def vec1_sub[symmetric] norm_vec1  vector_minus_component[symmetric] del: vector_minus_component)
+  apply (auto simp del: vector_minus_component)
   apply (erule_tac x=e in allE)
   apply clarify
   apply (rule_tac x=y in exI)
-  apply auto
+  apply (auto simp del: vector_minus_component)
   apply (rule order_le_less_trans)
   apply (rule component_le_norm)
   by auto
@@ -1450,7 +1449,7 @@
 text{* Uniqueness of the limit, when nontrivial. *}
 
 lemma Lim_unique:
-  fixes l::"real^'a" and net::"'b::zero_neq_one net"
+  fixes l::"real^'a::finite" and net::"'b::zero_neq_one net"
   assumes "\<not>(trivial_limit net)"  "(f ---> l) net"  "(f ---> l') net"
   shows "l = l'"
 proof-
@@ -1472,7 +1471,7 @@
 text{* Limit under bilinear function (surprisingly tedious, but important) *}
 
 lemma norm_bound_lemma:
-  "0 < e \<Longrightarrow> \<exists>d>0. \<forall>(x'::real^'b) y'::real^'a. norm(x' - (x::real^'b)) < d \<and> norm(y' - y) < d \<longrightarrow> norm(x') * norm(y' - y) + norm(x' - x) * norm(y) < e"
+  "0 < e \<Longrightarrow> \<exists>d>0. \<forall>(x'::real^'b::finite) y'::real^'a::finite. norm(x' - (x::real^'b)) < d \<and> norm(y' - y) < d \<longrightarrow> norm(x') * norm(y' - y) + norm(x' - x) * norm(y) < e"
 proof-
   assume e: "0 < e"
   have th1: "(2 * norm x + 2 * norm y + 2) > 0" using norm_ge_zero[of x] norm_ge_zero[of y] by norm
@@ -1494,8 +1493,7 @@
     have thy: "norm (y' - y) * norm x' < e / (2 * norm x + 2 * norm y + 2) * (1 + norm x)"
       using mult_strict_mono'[OF h(4) * norm_ge_zero norm_ge_zero] by auto
     also have "\<dots> \<le> e/2" apply simp unfolding divide_le_eq
-      using th1 th0 `e>0` apply auto
-      unfolding mult_assoc and real_mult_le_cancel_iff2[OF `e>0`] by auto
+      using th1 th0 `e>0` by auto
 
     finally have "norm x' * norm (y' - y) + norm (x' - x) * norm y < e"
       using thx and e by (simp add: field_simps)  }
@@ -1503,7 +1501,7 @@
 qed
 
 lemma Lim_bilinear:
-  fixes net :: "'a net" and h:: "real ^'m \<Rightarrow> real ^'n \<Rightarrow> real ^'p"
+  fixes net :: "'a net" and h:: "real ^'m::finite \<Rightarrow> real ^'n::finite \<Rightarrow> real ^'p::finite"
   assumes "(f ---> l) net" and "(g ---> m) net" and "bilinear h"
   shows "((\<lambda>x. h (f x) (g x)) ---> (h l m)) net"
 proof(cases "trivial_limit net")
@@ -1541,7 +1539,7 @@
 lemma Lim_at_id: "(id ---> a) (at a)"
 apply (subst within_UNIV[symmetric]) by (simp add: Lim_within_id)
 
-lemma Lim_at_zero: "(f ---> l) (at (a::real^'a)) \<longleftrightarrow> ((\<lambda>x. f(a + x)) ---> l) (at 0)" (is "?lhs = ?rhs")
+lemma Lim_at_zero: "(f ---> l) (at (a::real^'a::finite)) \<longleftrightarrow> ((\<lambda>x. f(a + x)) ---> l) (at 0)" (is "?lhs = ?rhs")
 proof
   assume "?lhs"
   { fix e::real assume "e>0"
@@ -1619,7 +1617,7 @@
 text{* Common case assuming being away from some crucial point like 0. *}
 
 lemma Lim_transform_away_within:
-  fixes f:: "real ^'m \<Rightarrow> real ^'n"
+  fixes f:: "real ^'m::finite \<Rightarrow> real ^'n::finite"
   assumes "a\<noteq>b" "\<forall>x\<in> S. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"
   and "(f ---> l) (at a within S)"
   shows "(g ---> l) (at a within S)"
@@ -1630,7 +1628,7 @@
 qed
 
 lemma Lim_transform_away_at:
-  fixes f:: "real ^'m \<Rightarrow> real ^'n"
+  fixes f:: "real ^'m::finite \<Rightarrow> real ^'n::finite"
   assumes ab: "a\<noteq>b" and fg: "\<forall>x. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"
   and fl: "(f ---> l) (at a)"
   shows "(g ---> l) (at a)"
@@ -1640,7 +1638,7 @@
 text{* Alternatively, within an open set. *}
 
 lemma Lim_transform_within_open:
-  fixes f:: "real ^'m \<Rightarrow> real ^'n"
+  fixes f:: "real ^'m::finite \<Rightarrow> real ^'n::finite"
   assumes "open S"  "a \<in> S"  "\<forall>x\<in>S. x \<noteq> a \<longrightarrow> f x = g x"  "(f ---> l) (at a)"
   shows "(g ---> l) (at a)"
 proof-
@@ -1917,7 +1915,7 @@
 subsection{* Boundedness. *}
 
   (* FIXME: This has to be unified with BSEQ!! *)
-definition "bounded S \<longleftrightarrow> (\<exists>a. \<forall>(x::real^'n) \<in> S. norm x <= a)"
+definition "bounded S \<longleftrightarrow> (\<exists>a. \<forall>(x::real^'n::finite) \<in> S. norm x <= a)"
 
 lemma bounded_empty[simp]: "bounded {}" by (simp add: bounded_def)
 lemma bounded_subset: "bounded T \<Longrightarrow> S \<subseteq> T ==> bounded S"
@@ -1978,7 +1976,7 @@
 lemma bounded_insert[intro]:"bounded(insert x S) \<longleftrightarrow> bounded S"
   by (metis Diff_cancel Un_empty_right Un_insert_right bounded_Un bounded_subset finite.emptyI finite_imp_bounded infinite_remove subset_insertI)
 
-lemma bot_bounded_UNIV[simp, intro]: "~(bounded (UNIV:: (real^'n) set))"
+lemma bot_bounded_UNIV[simp, intro]: "~(bounded (UNIV:: (real^'n::finite) set))"
 proof(auto simp add: bounded_pos not_le)
   fix b::real  assume b: "b >0"
   have b1: "b +1 \<ge> 0" using b by simp
@@ -1988,7 +1986,7 @@
 qed
 
 lemma bounded_linear_image:
-  fixes f :: "real^'m \<Rightarrow> real^'n"
+  fixes f :: "real^'m::finite \<Rightarrow> real^'n::finite"
   assumes "bounded S" "linear f"
   shows "bounded(f ` S)"
 proof-
@@ -2110,7 +2108,7 @@
 subsection{* Compactness (the definition is the one based on convegent subsequences). *}
 
 definition "compact S \<longleftrightarrow>
-   (\<forall>(f::nat \<Rightarrow> real^'n). (\<forall>n. f n \<in> S) \<longrightarrow>
+   (\<forall>(f::nat \<Rightarrow> real^'n::finite). (\<forall>n. f n \<in> S) \<longrightarrow>
        (\<exists>l\<in>S. \<exists>r. (\<forall>m n. m < n \<longrightarrow> r m < r n) \<and> ((f o r) ---> l) sequentially))"
 
 lemma monotone_bigger: fixes r::"nat\<Rightarrow>nat"
@@ -2178,81 +2176,69 @@
 qed
 
 lemma compact_lemma:
-  assumes "bounded s" and "\<forall>n. (x::nat \<Rightarrow>real^'a) n \<in> s"
-  shows "\<forall>d\<in>{1.. dimindex(UNIV::'a set)}.
-        \<exists>l::(real^'a). \<exists> r. (\<forall>n m::nat. m < n --> r m < r n) \<and>
-        (\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>i\<in>{1..d}. \<bar>x (r n) $ i - l $ i\<bar> < e)"
+  assumes "bounded s" and "\<forall>n. (x::nat \<Rightarrow>real^'a::finite) n \<in> s"
+  shows "\<forall>d.
+        \<exists>l::(real^'a::finite). \<exists> r. (\<forall>n m::nat. m < n --> r m < r n) \<and>
+        (\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>i\<in>d. \<bar>x (r n) $ i - l $ i\<bar> < e)"
 proof-
   obtain b where b:"\<forall>x\<in>s. norm x \<le> b" using assms(1)[unfolded bounded_def] by auto
-  { { fix i assume i:"i\<in>{1.. dimindex(UNIV::'a set)}"
+  { { fix i::'a
       { fix n::nat
-	have "\<bar>x n $ i\<bar> \<le> b" using b[THEN bspec[where x="x n"]] and component_le_norm[of i "x n"] and assms(2)[THEN spec[where x=n]] and i by auto }
+	have "\<bar>x n $ i\<bar> \<le> b" using b[THEN bspec[where x="x n"]] and component_le_norm[of "x n" i] and assms(2)[THEN spec[where x=n]] by auto }
       hence "\<forall>n. \<bar>x n $ i\<bar> \<le> b" by auto
     } note b' = this
 
-    fix d assume "d\<in>{1.. dimindex(UNIV::'a set)}"
+    fix d::"'a set" have "finite d" by simp
     hence "\<exists>l::(real^'a). \<exists> r. (\<forall>n m::nat. m < n --> r m < r n) \<and>
-        (\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>i\<in>{1..d}. \<bar>x (r n) $ i - l $ i\<bar> < e)"
-    proof(induct d) case 0 thus ?case by auto
-      (* The induction really starts at Suc 0 *)
-    next case (Suc d)
-      show ?case proof(cases "d = 0")
-	case True hence "Suc d = Suc 0" by auto
-	obtain l r where r:"\<forall>m n::nat. m < n \<longrightarrow> r m < r n" and lr:"\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<bar>x (r n) $ 1 - l\<bar> < e" using b' and dimindex_ge_1[of "UNIV::'a set"]
-	  using compact_real_lemma[of "\<lambda>i. (x i)$1" b] by auto
-	thus ?thesis apply(rule_tac x="vec l" in exI) apply(rule_tac x=r in exI)
-	  unfolding `Suc d = Suc 0` apply auto
-	  unfolding vec_component[OF Suc(2)[unfolded `Suc d = Suc 0`]] by auto
-      next
-	case False hence d:"d \<in>{1.. dimindex(UNIV::'a set)}" using Suc(2) by auto
-	obtain l1::"real^'a" and r1 where r1:"\<forall>n m::nat. m < n \<longrightarrow> r1 m < r1 n" and lr1:"\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>i\<in>{1..d}. \<bar>x (r1 n) $ i - l1 $ i\<bar> < e"
-	  using Suc(1)[OF d] by auto
-	obtain l2 r2 where r2:"\<forall>m n::nat. m < n \<longrightarrow> r2 m < r2 n" and lr2:"\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<bar>(x \<circ> r1) (r2 n) $ (Suc d) - l2\<bar> < e"
-	  using b'[OF Suc(2)] and compact_real_lemma[of "\<lambda>i. ((x \<circ> r1) i)$(Suc d)" b] by auto
+        (\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>i\<in>d. \<bar>x (r n) $ i - l $ i\<bar> < e)"
+    proof(induct d) case empty thus ?case by auto
+    next case (insert k d)
+	obtain l1::"real^'a" and r1 where r1:"\<forall>n m::nat. m < n \<longrightarrow> r1 m < r1 n" and lr1:"\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>i\<in>d. \<bar>x (r1 n) $ i - l1 $ i\<bar> < e"
+	  using insert(3) by auto
+	obtain l2 r2 where r2:"\<forall>m n::nat. m < n \<longrightarrow> r2 m < r2 n" and lr2:"\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<bar>(x \<circ> r1) (r2 n) $ k - l2\<bar> < e"
+	  using b'[of k] and compact_real_lemma[of "\<lambda>i. ((x \<circ> r1) i)$k" b] by auto
 	def r \<equiv> "r1 \<circ> r2" have r:"\<forall>m n. m < n \<longrightarrow> r m < r n" unfolding r_def o_def using r1 and r2 by auto
 	moreover
-	def l \<equiv> "(\<chi> i. if i = Suc d then l2 else l1$i)::real^'a"
+	def l \<equiv> "(\<chi> i. if i = k then l2 else l1$i)::real^'a"
 	{ fix e::real assume "e>0"
-	  from lr1 obtain N1 where N1:"\<forall>n\<ge>N1. \<forall>i\<in>{1..d}. \<bar>x (r1 n) $ i - l1 $ i\<bar> < e" using `e>0` by blast
-	  from lr2 obtain N2 where N2:"\<forall>n\<ge>N2. \<bar>(x \<circ> r1) (r2 n) $ (Suc d) - l2\<bar> < e" using `e>0` by blast
+	  from lr1 obtain N1 where N1:"\<forall>n\<ge>N1. \<forall>i\<in>d. \<bar>x (r1 n) $ i - l1 $ i\<bar> < e" using `e>0` by blast
+	  from lr2 obtain N2 where N2:"\<forall>n\<ge>N2. \<bar>(x \<circ> r1) (r2 n) $ k - l2\<bar> < e" using `e>0` by blast
 	  { fix n assume n:"n\<ge> N1 + N2"
-	    fix i assume i:"i\<in>{1..Suc d}" hence i':"i\<in>{1.. dimindex(UNIV::'a set)}" using Suc by auto
+	    fix i assume i:"i\<in>(insert k d)"
 	    hence "\<bar>x (r n) $ i - l $ i\<bar> < e"
 	      using N2[THEN spec[where x="n"]] and n
  	      using N1[THEN spec[where x="r2 n"]] and n
 	      using monotone_bigger[OF r] and i
-	      unfolding l_def and r_def and Cart_lambda_beta'[OF i']
+	      unfolding l_def and r_def
 	      using monotone_bigger[OF r2, of n] by auto  }
-	  hence "\<exists>N. \<forall>n\<ge>N. \<forall>i\<in>{1..Suc d}. \<bar>x (r n) $ i - l $ i\<bar> < e" by blast	}
-	ultimately show ?thesis by auto
-      qed
+	  hence "\<exists>N. \<forall>n\<ge>N. \<forall>i\<in>(insert k d). \<bar>x (r n) $ i - l $ i\<bar> < e" by blast	}
+	ultimately show ?case by auto
     qed  }
   thus ?thesis by auto
 qed
 
-lemma bounded_closed_imp_compact: fixes s::"(real^'a) set"
+lemma bounded_closed_imp_compact: fixes s::"(real^'a::finite) set"
   assumes "bounded s" and "closed s"
   shows "compact s"
 proof-
-  let ?d = "dimindex (UNIV::'a set)"
+  let ?d = "UNIV::'a set"
   { fix f assume as:"\<forall>n::nat. f n \<in> s"
     obtain l::"real^'a" and r where r:"\<forall>n m::nat. m < n \<longrightarrow> r m < r n"
-      and lr:"\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>i\<in>{1..?d}. \<bar>f (r n) $ i - l $ i\<bar> < e"
-      using compact_lemma[OF assms(1) as, THEN bspec[where x="?d"]] and dimindex_ge_1[of "UNIV::'a set"] by auto
+      and lr:"\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>i\<in>?d. \<bar>f (r n) $ i - l $ i\<bar> < e"
+      using compact_lemma[OF assms(1) as, THEN spec[where x="?d"]] by auto
     { fix e::real assume "e>0"
-      hence "0 < e / (real_of_nat ?d)" using dimindex_nonzero[of "UNIV::'a set"] using divide_pos_pos[of e, of "real_of_nat ?d"] by auto
-      then obtain N::nat where N:"\<forall>n\<ge>N. \<forall>i\<in>{1..?d}. \<bar>f (r n) $ i - l $ i\<bar> < e / (real_of_nat ?d)" using lr[THEN spec[where x="e / (real_of_nat ?d)"]] by blast
+      hence "0 < e / (real_of_nat (card ?d))" using zero_less_card_finite using divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto
+      then obtain N::nat where N:"\<forall>n\<ge>N. \<forall>i\<in>?d. \<bar>f (r n) $ i - l $ i\<bar> < e / (real_of_nat (card ?d))" using lr[THEN spec[where x="e / (real_of_nat (card ?d))"]] by blast
       { fix n assume n:"n\<ge>N"
-	have "1 \<in> {1..?d}" using dimindex_nonzero[of "UNIV::'a set"] by auto
-	hence "finite {1..?d}"  "{1..?d} \<noteq> {}" by auto
+	hence "finite ?d"  "?d \<noteq> {}" by auto
 	moreover
-	{ fix i assume i:"i \<in> {1..?d}"
-	  hence "\<bar>((f \<circ> r) n - l) $ i\<bar> < e / real_of_nat ?d" using `n\<ge>N` using N[THEN spec[where x=n]]
-	    apply auto apply(erule_tac x=i in ballE) unfolding vector_minus_component[OF i] by auto  }
-	ultimately have "(\<Sum>i = 1..?d. \<bar>((f \<circ> r) n - l) $ i\<bar>)
-	  < (\<Sum>i = 1..?d. e / real_of_nat ?d)"
-	  using setsum_strict_mono[of "{1..?d}" "\<lambda>i. \<bar>((f \<circ> r) n - l) $ i\<bar>" "\<lambda>i. e / (real_of_nat ?d)"] by auto
-	hence "(\<Sum>i = 1..?d. \<bar>((f \<circ> r) n - l) $ i\<bar>) < e" unfolding setsum_constant using dimindex_nonzero[of "UNIV::'a set"] by auto
+	{ fix i assume i:"i \<in> ?d"
+	  hence "\<bar>((f \<circ> r) n - l) $ i\<bar> < e / real_of_nat (card ?d)" using `n\<ge>N` using N[THEN spec[where x=n]]
+	    by auto  }
+	ultimately have "(\<Sum>i \<in> ?d. \<bar>((f \<circ> r) n - l) $ i\<bar>)
+	  < (\<Sum>i \<in> ?d. e / real_of_nat (card ?d))"
+	  using setsum_strict_mono[of "?d" "\<lambda>i. \<bar>((f \<circ> r) n - l) $ i\<bar>" "\<lambda>i. e / (real_of_nat (card ?d))"] by auto
+	hence "(\<Sum>i \<in> ?d. \<bar>((f \<circ> r) n - l) $ i\<bar>) < e" unfolding setsum_constant by auto
 	hence "dist ((f \<circ> r) n) l < e" unfolding dist_def using norm_le_l1[of "(f \<circ> r) n - l"] by auto  }
       hence "\<exists>N. \<forall>n\<ge>N. dist ((f \<circ> r) n) l < e" by auto  }
     hence *:"((f \<circ> r) ---> l) sequentially" unfolding Lim_sequentially by auto
@@ -2268,7 +2254,7 @@
 
 definition cauchy_def:"cauchy s \<longleftrightarrow> (\<forall>e>0. \<exists>N. \<forall>m n. m \<ge> N \<and> n \<ge> N --> dist(s m)(s n) < e)"
 
-definition complete_def:"complete s \<longleftrightarrow> (\<forall>f::(nat=>real^'a). (\<forall>n. f n \<in> s) \<and> cauchy f
+definition complete_def:"complete s \<longleftrightarrow> (\<forall>f::(nat=>real^'a::finite). (\<forall>n. f n \<in> s) \<and> cauchy f
                       --> (\<exists>l \<in> s. (f ---> l) sequentially))"
 
 lemma cauchy: "cauchy s \<longleftrightarrow> (\<forall>e>0.\<exists> N::nat. \<forall>n\<ge>N. dist(s n)(s N) < e)" (is "?lhs = ?rhs")
@@ -2350,7 +2336,7 @@
 lemma complete_univ:
  "complete UNIV"
 proof(simp add: complete_def, rule, rule)
-  fix f::"nat \<Rightarrow> real^'n" assume "cauchy f"
+  fix f::"nat \<Rightarrow> real^'n::finite" assume "cauchy f"
   hence "bounded (f`UNIV)" using cauchy_imp_bounded[of f] unfolding image_def by auto
   hence "compact (closure (f`UNIV))"  using bounded_closed_imp_compact[of "closure (range f)"] by auto
   hence "complete (closure (range f))" using compact_imp_complete by auto
@@ -2389,7 +2375,7 @@
 
 subsection{* Total boundedness. *}
 
-fun helper_1::"((real^'n) set) \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real^'n" where
+fun helper_1::"((real^'n::finite) set) \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real^'n" where
   "helper_1 s e n = (SOME y::real^'n. y \<in> s \<and> (\<forall>m<n. \<not> (dist (helper_1 s e m) y < e)))"
 declare helper_1.simps[simp del]
 
@@ -2422,7 +2408,7 @@
 
 subsection{* Heine-Borel theorem (following Burkill \& Burkill vol. 2) *}
 
-lemma heine_borel_lemma: fixes s::"(real^'n) set"
+lemma heine_borel_lemma: fixes s::"(real^'n::finite) set"
   assumes "compact s"  "s \<subseteq> (\<Union> t)"  "\<forall>b \<in> t. open b"
   shows "\<exists>e>0. \<forall>x \<in> s. \<exists>b \<in> t. ball x e \<subseteq> b"
 proof(rule ccontr)
@@ -2513,11 +2499,11 @@
 
 subsection{* Complete the chain of compactness variants. *}
 
-primrec helper_2::"(real \<Rightarrow> real^'n) \<Rightarrow> nat \<Rightarrow> real ^'n" where
+primrec helper_2::"(real \<Rightarrow> real^'n::finite) \<Rightarrow> nat \<Rightarrow> real ^'n" where
   "helper_2 beyond 0 = beyond 0" |
   "helper_2 beyond (Suc n) = beyond (norm (helper_2 beyond n) + 1 )"
 
-lemma bolzano_weierstrass_imp_bounded: fixes s::"(real^'n) set"
+lemma bolzano_weierstrass_imp_bounded: fixes s::"(real^'n::finite) set"
   assumes "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t)"
   shows "bounded s"
 proof(rule ccontr)
@@ -2576,7 +2562,7 @@
 
 lemma sequence_infinite_lemma:
   assumes "\<forall>n::nat. (f n  \<noteq> l)"  "(f ---> l) sequentially"
-  shows "infinite {y::real^'a. (\<exists> n. y = f n)}"
+  shows "infinite {y::real^'a::finite. (\<exists> n. y = f n)}"
 proof(rule ccontr)
   let ?A = "(\<lambda>x. dist x l) ` {y. \<exists>n. y = f n}"
   assume "\<not> infinite {y. \<exists>n. y = f n}"
@@ -2771,7 +2757,7 @@
 lemma bounded_closed_nest:
   assumes "\<forall>n. closed(s n)" "\<forall>n. (s n \<noteq> {})"
   "(\<forall>m n. m \<le> n --> s n \<subseteq> s m)"  "bounded(s 0)"
-  shows "\<exists> a::real^'a. \<forall>n::nat. a \<in> s(n)"
+  shows "\<exists> a::real^'a::finite. \<forall>n::nat. a \<in> s(n)"
 proof-
   from assms(2) obtain x where x:"\<forall>n::nat. x n \<in> s n" using choice[of "\<lambda>n x. x\<in> s n"] by auto
   from assms(4,1) have *:"compact (s 0)" using bounded_closed_imp_compact[of "s 0"] by auto
@@ -2803,7 +2789,7 @@
           "\<forall>n. (s n \<noteq> {})"
           "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
           "\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y \<in> (s n). dist x y < e"
-  shows "\<exists>a::real^'a. \<forall>n::nat. a \<in> s n"
+  shows "\<exists>a::real^'a::finite. \<forall>n::nat. a \<in> s n"
 proof-
   have "\<forall>n. \<exists> x. x\<in>s n" using assms(2) by auto
   hence "\<exists>t. \<forall>n. t n \<in> s n" using choice[of "\<lambda> n x. x \<in> s n"] by auto
@@ -2836,7 +2822,7 @@
           "\<forall>n. s n \<noteq> {}"
           "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
           "\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y\<in>(s n). dist x y < e"
-  shows "\<exists>a::real^'a. \<Inter> {t. (\<exists>n::nat. t = s n)} = {a}"
+  shows "\<exists>a::real^'a::finite. \<Inter> {t. (\<exists>n::nat. t = s n)} = {a}"
 proof-
   obtain a where a:"\<forall>n. a \<in> s n" using decreasing_closed_nest[of s] using assms by auto
   { fix b assume b:"b \<in> \<Inter>{t. \<exists>n. t = s n}"
@@ -2851,7 +2837,7 @@
 
 text{* Cauchy-type criteria for uniform convergence. *}
 
-lemma uniformly_convergent_eq_cauchy: fixes s::"nat \<Rightarrow> 'b \<Rightarrow> real^'a" shows
+lemma uniformly_convergent_eq_cauchy: fixes s::"nat \<Rightarrow> 'b \<Rightarrow> real^'a::finite" shows
  "(\<exists>l. \<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x --> dist(s n x)(l x) < e) \<longleftrightarrow>
   (\<forall>e>0. \<exists>N. \<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x  --> dist (s m x) (s n x) < e)" (is "?lhs = ?rhs")
 proof(rule)
@@ -2960,7 +2946,7 @@
     apply (auto simp add: dist_sym) apply(erule_tac x=e in allE) by auto
 qed
 
-lemma continuous_at_ball: fixes f::"real^'a \<Rightarrow> real^'a"
+lemma continuous_at_ball: fixes f::"real^'a::finite \<Rightarrow> real^'a"
   shows "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. f ` (ball x d) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")
 proof
   assume ?lhs thus ?rhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball
@@ -3255,7 +3241,7 @@
 
 lemma uniformly_continuous_on_add:
   assumes "uniformly_continuous_on s f" "uniformly_continuous_on s g"
-  shows "uniformly_continuous_on s (\<lambda>x. f(x) + g(x) ::real^'n)"
+  shows "uniformly_continuous_on s (\<lambda>x. f(x) + g(x) ::real^'n::finite)"
 proof-
   have *:"\<And>fx fy gx gy::real^'n. fx - fy + (gx - gy) = fx + gx - (fy + gy)" by auto
   {  fix x y assume "((\<lambda>n. f (x n) - f (y n)) ---> 0) sequentially"
@@ -3570,7 +3556,7 @@
     { fix y assume "dist y (c *s x) < e * \<bar>c\<bar>"
       hence "norm ((1 / c) *s y - x) < e" unfolding dist_def
 	using norm_mul[of c "(1 / c) *s y - x", unfolded vector_ssub_ldistrib, unfolded vector_smult_assoc] assms(1)
-	  mult_less_imp_less_left[of "abs c" "norm ((1 / c) *s y - x)" e, unfolded real_mult_commute[of "abs c" e]] assms(1)[unfolded zero_less_abs_iff[THEN sym]] by simp
+	  assms(1)[unfolded zero_less_abs_iff[THEN sym]] by (simp del:zero_less_abs_iff)
       hence "y \<in> op *s c ` s" using rev_image_eqI[of "(1 / c) *s y" s y "op *s c"]  e[THEN spec[where x="(1 / c) *s y"]]  assms(1) unfolding dist_def vector_smult_assoc by auto  }
     ultimately have "\<exists>e>0. \<forall>x'. dist x' (c *s x) < e \<longrightarrow> x' \<in> op *s c ` s" apply(rule_tac x="e * abs c" in exI) by auto  }
   thus ?thesis unfolding open_def by auto
@@ -3729,7 +3715,7 @@
 
 subsection{* Topological properties of linear functions.                               *}
 
-lemma linear_lim_0: fixes f::"real^'a \<Rightarrow> real^'b"
+lemma linear_lim_0: fixes f::"real^'a::finite \<Rightarrow> real^'b::finite"
   assumes "linear f" shows "(f ---> 0) (at (0))"
 proof-
   obtain B where "B>0" and B:"\<forall>x. norm (f x) \<le> B * norm x" using linear_bounded_pos[OF assms] by auto
@@ -3813,19 +3799,18 @@
 unfolding continuous_on_vec1_range norm_vec1[THEN sym] by (metis norm_vec1 order_le_less_trans real_abs_sub_norm)
 
 lemma continuous_at_vec1_component:
-  assumes "1 \<le> i" "i \<le> dimindex(UNIV::('a set))"
-  shows "continuous (at (a::real^'a)) (\<lambda> x. vec1(x$i))"
+  shows "continuous (at (a::real^'a::finite)) (\<lambda> x. vec1(x$i))"
 proof-
   { fix e::real and x assume "0 < dist x a" "dist x a < e" "e>0"
-    hence "\<bar>x $ i - a $ i\<bar> < e" using component_le_norm[of i "x - a"] vector_minus_component[of i x a] assms unfolding dist_def by auto  }
+    hence "\<bar>x $ i - a $ i\<bar> < e" using component_le_norm[of "x - a" i] unfolding dist_def by auto  }
   thus ?thesis unfolding continuous_at tendsto_def eventually_at dist_vec1 by auto
 qed
 
 lemma continuous_on_vec1_component:
-  assumes "i \<in> {1..dimindex (UNIV::'a set)}"  shows "continuous_on s (\<lambda> x::real^'a. vec1(x$i))"
+  shows "continuous_on s (\<lambda> x::real^'a::finite. vec1(x$i))"
 proof-
   { fix e::real and x xa assume "x\<in>s" "e>0" "xa\<in>s" "0 < norm (xa - x) \<and> norm (xa - x) < e"
-    hence "\<bar>xa $ i - x $ i\<bar> < e" using component_le_norm[of i "xa - x"] vector_minus_component[of i xa x] assms by auto  }
+    hence "\<bar>xa $ i - x $ i\<bar> < e" using component_le_norm[of "xa - x" i] by auto  }
   thus ?thesis unfolding continuous_on Lim_within dist_vec1 unfolding dist_def by auto
 qed
 
@@ -3970,7 +3955,7 @@
       hence fx0:"f x \<noteq> 0" using `l \<noteq> 0` by auto
       hence fxl0: "(f x) * l \<noteq> 0" using `l \<noteq> 0` by auto
       from * have **:"\<bar>f x - l\<bar> < l\<twosuperior> * e / 2" by auto
-      have "\<bar>f x\<bar> * 2 \<ge> \<bar>l\<bar>" using * by (auto simp del: Arith_Tools.less_divide_eq_number_of1)
+      have "\<bar>f x\<bar> * 2 \<ge> \<bar>l\<bar>" using * by (auto simp del: less_divide_eq_number_of1)
       hence "\<bar>f x\<bar> * 2 * \<bar>l\<bar>  \<ge> \<bar>l\<bar> * \<bar>l\<bar>" unfolding mult_le_cancel_right by auto
       hence "\<bar>f x * l\<bar> * 2  \<ge> \<bar>l\<bar>^2" unfolding real_mult_commute and power2_eq_square by auto
       hence ***:"inverse \<bar>f x * l\<bar> \<le> inverse (l\<twosuperior> / 2)" using fxl0
@@ -4070,7 +4055,7 @@
 proof-
   have *:"{x + y | x y. x \<in> s \<and> y \<in> t} =(\<lambda>z. fstcart z + sndcart z) ` {pastecart x y | x y.  x \<in> s \<and> y \<in> t}"
     apply auto unfolding image_iff apply(rule_tac x="pastecart xa y" in bexI) unfolding fstcart_pastecart sndcart_pastecart by auto
-  have "linear (\<lambda>z::real^('a, 'a) finite_sum. fstcart z + sndcart z)" unfolding linear_def
+  have "linear (\<lambda>z::real^('a + 'a). fstcart z + sndcart z)" unfolding linear_def
     unfolding fstcart_add sndcart_add apply auto
     unfolding vector_add_ldistrib fstcart_cmul[THEN sym] sndcart_cmul[THEN sym] by auto
   hence "continuous_on {pastecart x y |x y. x \<in> s \<and> y \<in> t} (\<lambda>z. fstcart z + sndcart z)"
@@ -4306,90 +4291,86 @@
 
 (* A cute way of denoting open and closed intervals using overloading.       *)
 
-lemma interval: fixes a :: "'a::ord^'n" shows
-  "{a <..< b} = {x::'a^'n. \<forall>i \<in> dimset a. a$i < x$i \<and> x$i < b$i}" and
-  "{a .. b} = {x::'a^'n. \<forall>i \<in> dimset a. a$i \<le> x$i \<and> x$i \<le> b$i}"
+lemma interval: fixes a :: "'a::ord^'n::finite" shows
+  "{a <..< b} = {x::'a^'n. \<forall>i. a$i < x$i \<and> x$i < b$i}" and
+  "{a .. b} = {x::'a^'n. \<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i}"
   by (auto simp add: expand_set_eq vector_less_def vector_less_eq_def)
 
-lemma mem_interval:
-  "x \<in> {a<..<b} \<longleftrightarrow> (\<forall>i \<in> dimset a. a$i < x$i \<and> x$i < b$i)"
-  "x \<in> {a .. b} \<longleftrightarrow> (\<forall>i \<in> dimset a. a$i \<le> x$i \<and> x$i \<le> b$i)"
+lemma mem_interval: fixes a :: "'a::ord^'n::finite" shows
+  "x \<in> {a<..<b} \<longleftrightarrow> (\<forall>i. a$i < x$i \<and> x$i < b$i)"
+  "x \<in> {a .. b} \<longleftrightarrow> (\<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i)"
   using interval[of a b]
   by(auto simp add: expand_set_eq vector_less_def vector_less_eq_def)
 
-lemma interval_eq_empty: fixes a :: "real^'n" shows
- "({a <..< b} = {} \<longleftrightarrow> (\<exists>i \<in> dimset a. b$i \<le> a$i))" (is ?th1) and
- "({a  ..  b} = {} \<longleftrightarrow> (\<exists>i \<in> dimset a. b$i < a$i))" (is ?th2)
+lemma interval_eq_empty: fixes a :: "real^'n::finite" shows
+ "({a <..< b} = {} \<longleftrightarrow> (\<exists>i. b$i \<le> a$i))" (is ?th1) and
+ "({a  ..  b} = {} \<longleftrightarrow> (\<exists>i. b$i < a$i))" (is ?th2)
 proof-
-  { fix i x assume i:"i\<in>dimset a" and as:"b$i \<le> a$i" and x:"x\<in>{a <..< b}"
+  { fix i x assume as:"b$i \<le> a$i" and x:"x\<in>{a <..< b}"
     hence "a $ i < x $ i \<and> x $ i < b $ i" unfolding mem_interval by auto
     hence "a$i < b$i" by auto
     hence False using as by auto  }
   moreover
-  { assume as:"\<forall>i \<in> dimset a. \<not> (b$i \<le> a$i)"
+  { assume as:"\<forall>i. \<not> (b$i \<le> a$i)"
     let ?x = "(1/2) *s (a + b)"
-    { fix i assume i:"i\<in>dimset a"
-      hence "a$i < b$i" using as[THEN bspec[where x=i]] by auto
+    { fix i
+      have "a$i < b$i" using as[THEN spec[where x=i]] by auto
       hence "a$i < ((1/2) *s (a+b)) $ i" "((1/2) *s (a+b)) $ i < b$i"
-	unfolding vector_smult_component[OF i] and vector_add_component[OF i]
-	by (auto simp add: Arith_Tools.less_divide_eq_number_of1)  }
+	unfolding vector_smult_component and vector_add_component
+	by (auto simp add: less_divide_eq_number_of1)  }
     hence "{a <..< b} \<noteq> {}" using mem_interval(1)[of "?x" a b] by auto  }
   ultimately show ?th1 by blast
 
-  { fix i x assume i:"i\<in>dimset a" and as:"b$i < a$i" and x:"x\<in>{a .. b}"
+  { fix i x assume as:"b$i < a$i" and x:"x\<in>{a .. b}"
     hence "a $ i \<le> x $ i \<and> x $ i \<le> b $ i" unfolding mem_interval by auto
     hence "a$i \<le> b$i" by auto
     hence False using as by auto  }
   moreover
-  { assume as:"\<forall>i \<in> dimset a. \<not> (b$i < a$i)"
+  { assume as:"\<forall>i. \<not> (b$i < a$i)"
     let ?x = "(1/2) *s (a + b)"
-    { fix i assume i:"i\<in>dimset a"
-      hence "a$i \<le> b$i" using as[THEN bspec[where x=i]] by auto
+    { fix i
+      have "a$i \<le> b$i" using as[THEN spec[where x=i]] by auto
       hence "a$i \<le> ((1/2) *s (a+b)) $ i" "((1/2) *s (a+b)) $ i \<le> b$i"
-	unfolding vector_smult_component[OF i] and vector_add_component[OF i]
-	by (auto simp add: Arith_Tools.less_divide_eq_number_of1)  }
+	unfolding vector_smult_component and vector_add_component
+	by (auto simp add: less_divide_eq_number_of1)  }
     hence "{a .. b} \<noteq> {}" using mem_interval(2)[of "?x" a b] by auto  }
   ultimately show ?th2 by blast
 qed
 
-lemma interval_ne_empty: fixes a :: "real^'n" shows
-  "{a  ..  b} \<noteq> {} \<longleftrightarrow> (\<forall>i \<in> dimset a. a$i \<le> b$i)" and
-  "{a <..< b} \<noteq> {} \<longleftrightarrow> (\<forall>i \<in> dimset a. a$i < b$i)"
-  unfolding interval_eq_empty[of a b] by auto
-
-lemma subset_interval_imp: fixes a :: "real^'n" shows
- "(\<forall>i \<in> dimset a. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c .. d} \<subseteq> {a .. b}" and
- "(\<forall>i \<in> dimset a. a$i < c$i \<and> d$i < b$i) \<Longrightarrow> {c .. d} \<subseteq> {a<..<b}" and
- "(\<forall>i \<in> dimset a. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c<..<d} \<subseteq> {a .. b}" and
- "(\<forall>i \<in> dimset a. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c<..<d} \<subseteq> {a<..<b}"
-  unfolding subset_eq[unfolded Ball_def] unfolding mem_interval by(auto elim!: ballE)
-
-lemma interval_sing: fixes a :: "'a::linorder^'n" shows
+lemma interval_ne_empty: fixes a :: "real^'n::finite" shows
+  "{a  ..  b} \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i \<le> b$i)" and
+  "{a <..< b} \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i < b$i)"
+  unfolding interval_eq_empty[of a b] by (auto simp add: not_less not_le) (* BH: Why doesn't just "auto" work here? *)
+
+lemma subset_interval_imp: fixes a :: "real^'n::finite" shows
+ "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c .. d} \<subseteq> {a .. b}" and
+ "(\<forall>i. a$i < c$i \<and> d$i < b$i) \<Longrightarrow> {c .. d} \<subseteq> {a<..<b}" and
+ "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c<..<d} \<subseteq> {a .. b}" and
+ "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c<..<d} \<subseteq> {a<..<b}"
+  unfolding subset_eq[unfolded Ball_def] unfolding mem_interval
+  by (auto intro: order_trans less_le_trans le_less_trans less_imp_le) (* BH: Why doesn't just "auto" work here? *)
+
+lemma interval_sing: fixes a :: "'a::linorder^'n::finite" shows
  "{a .. a} = {a} \<and> {a<..<a} = {}"
 apply(auto simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
-apply (simp only: order_eq_iff)
-using dimindex_ge_1[of "UNIV :: 'n set"]
-apply (auto simp add: not_less )
-apply (erule_tac x= 1 in ballE)
-apply (rule bexI[where x=1])
-apply auto
+apply (simp add: order_eq_iff)
+apply (auto simp add: not_less less_imp_le)
 done
 
 
-lemma interval_open_subset_closed:  fixes a :: "'a::preorder^'n" shows
+lemma interval_open_subset_closed:  fixes a :: "'a::preorder^'n::finite" shows
  "{a<..<b} \<subseteq> {a .. b}"
 proof(simp add: subset_eq, rule)
   fix x
   assume x:"x \<in>{a<..<b}"
-  { fix i assume "i \<in> dimset a"
-    hence "a $ i \<le> x $ i"
+  { fix i
+    have "a $ i \<le> x $ i"
       using x order_less_imp_le[of "a$i" "x$i"]
       by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
   }
   moreover
-  { fix i assume "i \<in> dimset a"
-    hence "x $ i \<le> b $ i"
-      using x
+  { fix i
+    have "x $ i \<le> b $ i"
       using x order_less_imp_le[of "x$i" "b$i"]
       by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
   }
@@ -4398,76 +4379,76 @@
     by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
 qed
 
-lemma subset_interval: fixes a :: "real^'n" shows
- "{c .. d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i \<in> dimset a. c$i \<le> d$i) --> (\<forall>i \<in> dimset a. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th1) and
- "{c .. d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i \<in> dimset a. c$i \<le> d$i) --> (\<forall>i \<in> dimset a. a$i < c$i \<and> d$i < b$i)" (is ?th2) and
- "{c<..<d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i \<in> dimset a. c$i < d$i) --> (\<forall>i \<in> dimset a. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th3) and
- "{c<..<d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i \<in> dimset a. c$i < d$i) --> (\<forall>i \<in> dimset a. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th4)
+lemma subset_interval: fixes a :: "real^'n::finite" shows
+ "{c .. d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th1) and
+ "{c .. d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i < c$i \<and> d$i < b$i)" (is ?th2) and
+ "{c<..<d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th3) and
+ "{c<..<d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th4)
 proof-
-  show ?th1 unfolding subset_eq and Ball_def and mem_interval apply auto by(erule_tac x=xa in allE, simp)+
-  show ?th2 unfolding subset_eq and Ball_def and mem_interval apply auto by(erule_tac x=xa in allE, simp)+
-  { assume as: "{c<..<d} \<subseteq> {a .. b}" "\<forall>i \<in> dimset a. c$i < d$i"
-    hence "{c<..<d} \<noteq> {}" unfolding interval_eq_empty by auto
-    fix i assume i:"i \<in> dimset a"
+  show ?th1 unfolding subset_eq and Ball_def and mem_interval by (auto intro: order_trans)
+  show ?th2 unfolding subset_eq and Ball_def and mem_interval by (auto intro: le_less_trans less_le_trans order_trans less_imp_le)
+  { assume as: "{c<..<d} \<subseteq> {a .. b}" "\<forall>i. c$i < d$i"
+    hence "{c<..<d} \<noteq> {}" unfolding interval_eq_empty by (auto, drule_tac x=i in spec, simp) (* BH: Why doesn't just "auto" work? *)
+    fix i
     (** TODO combine the following two parts as done in the HOL_light version. **)
     { let ?x = "(\<chi> j. (if j=i then ((min (a$j) (d$j))+c$j)/2 else (c$j+d$j)/2))::real^'n"
       assume as2: "a$i > c$i"
-      { fix j assume j:"j\<in>dimset a"
-	hence "c $ j < ?x $ j \<and> ?x $ j < d $ j" unfolding Cart_lambda_beta[THEN bspec[where x=j], OF j]
-	  apply(cases "j=i") using as(2)[THEN bspec[where x=j], OF j]
-	  by (auto simp add: Arith_Tools.less_divide_eq_number_of1 as2)  }
+      { fix j
+	have "c $ j < ?x $ j \<and> ?x $ j < d $ j" unfolding Cart_lambda_beta
+	  apply(cases "j=i") using as(2)[THEN spec[where x=j]]
+	  by (auto simp add: less_divide_eq_number_of1 as2)  }
       hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
       moreover
       have "?x\<notin>{a .. b}"
-	unfolding mem_interval apply auto apply(rule_tac x=i in bexI)
-	unfolding Cart_lambda_beta[THEN bspec[where x=i], OF i]
-	using as(2)[THEN bspec[where x=i], OF i] and as2 and i
-	by (auto simp add: Arith_Tools.less_divide_eq_number_of1)
+	unfolding mem_interval apply auto apply(rule_tac x=i in exI)
+	using as(2)[THEN spec[where x=i]] and as2
+	by (auto simp add: less_divide_eq_number_of1)
       ultimately have False using as by auto  }
     hence "a$i \<le> c$i" by(rule ccontr)auto
     moreover
     { let ?x = "(\<chi> j. (if j=i then ((max (b$j) (c$j))+d$j)/2 else (c$j+d$j)/2))::real^'n"
       assume as2: "b$i < d$i"
-      { fix j assume j:"j\<in>dimset a"
-	hence "d $ j > ?x $ j \<and> ?x $ j > c $ j" unfolding Cart_lambda_beta[THEN bspec[where x=j], OF j]
-	  apply(cases "j=i") using as(2)[THEN bspec[where x=j], OF j]
-	  by (auto simp add: Arith_Tools.less_divide_eq_number_of1 as2)  }
+      { fix j
+	have "d $ j > ?x $ j \<and> ?x $ j > c $ j" unfolding Cart_lambda_beta
+	  apply(cases "j=i") using as(2)[THEN spec[where x=j]]
+	  by (auto simp add: less_divide_eq_number_of1 as2)  }
       hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
       moreover
       have "?x\<notin>{a .. b}"
-	unfolding mem_interval apply auto apply(rule_tac x=i in bexI)
-	unfolding Cart_lambda_beta[THEN bspec[where x=i], OF i]
-	using as(2)[THEN bspec[where x=i], OF i] and as2 and i
-	by (auto simp add: Arith_Tools.less_divide_eq_number_of1)
+	unfolding mem_interval apply auto apply(rule_tac x=i in exI)
+	using as(2)[THEN spec[where x=i]] and as2
+	by (auto simp add: less_divide_eq_number_of1)
       ultimately have False using as by auto  }
     hence "b$i \<ge> d$i" by(rule ccontr)auto
     ultimately
     have "a$i \<le> c$i \<and> d$i \<le> b$i" by auto
   } note part1 = this
-  thus ?th3 unfolding subset_eq and Ball_def and mem_interval apply auto by(erule_tac x=xa in allE, simp)+
-  { assume as:"{c<..<d} \<subseteq> {a<..<b}" "\<forall>i \<in> dimset a. c$i < d$i"
-    fix i assume i:"i \<in> dimset a"
+  thus ?th3 unfolding subset_eq and Ball_def and mem_interval apply auto apply (erule_tac x=ia in allE, simp)+ by (erule_tac x=i in allE, erule_tac x=i in allE, simp)+
+  { assume as:"{c<..<d} \<subseteq> {a<..<b}" "\<forall>i. c$i < d$i"
+    fix i
     from as(1) have "{c<..<d} \<subseteq> {a..b}" using interval_open_subset_closed[of a b] by auto
-    hence "a$i \<le> c$i \<and> d$i \<le> b$i" using part1 and as(2) and i by auto  } note * = this
-  thus ?th4 unfolding subset_eq and Ball_def and mem_interval apply auto by(erule_tac x=xa in allE, simp)+
-qed
-
-lemma disjoint_interval: fixes a::"real^'n" shows
-  "{a .. b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i \<in> dimset a. (b$i < a$i \<or> d$i < c$i \<or> b$i < c$i \<or> d$i < a$i))" (is ?th1) and
-  "{a .. b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i \<in> dimset a. (b$i < a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th2) and
-  "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i \<in> dimset a. (b$i \<le> a$i \<or> d$i < c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th3) and
-  "{a<..<b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i \<in> dimset a. (b$i \<le> a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th4)
+    hence "a$i \<le> c$i \<and> d$i \<le> b$i" using part1 and as(2) by auto  } note * = this
+  thus ?th4 unfolding subset_eq and Ball_def and mem_interval apply auto apply (erule_tac x=ia in allE, simp)+ by (erule_tac x=i in allE, erule_tac x=i in allE, simp)+
+qed
+
+lemma disjoint_interval: fixes a::"real^'n::finite" shows
+  "{a .. b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i < c$i \<or> b$i < c$i \<or> d$i < a$i))" (is ?th1) and
+  "{a .. b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th2) and
+  "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i < c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th3) and
+  "{a<..<b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th4)
 proof-
   let ?z = "(\<chi> i. ((max (a$i) (c$i)) + (min (b$i) (d$i))) / 2)::real^'n"
   show ?th1 ?th2 ?th3 ?th4
-  unfolding expand_set_eq and Int_iff and empty_iff and mem_interval and ball_conj_distrib[THEN sym] and eq_False
-  by (auto simp add: Cart_lambda_beta' Arith_Tools.less_divide_eq_number_of1 intro!: bexI elim!: allE[where x="?z"])
-qed
-
-lemma inter_interval: fixes a :: "'a::linorder^'n" shows
+  unfolding expand_set_eq and Int_iff and empty_iff and mem_interval and all_conj_distrib[THEN sym] and eq_False
+  apply (auto elim!: allE[where x="?z"])
+  apply ((rule_tac x=x in exI, force) | (rule_tac x=i in exI, force))+
+  done
+qed
+
+lemma inter_interval: fixes a :: "'a::linorder^'n::finite" shows
  "{a .. b} \<inter> {c .. d} =  {(\<chi> i. max (a$i) (c$i)) .. (\<chi> i. min (b$i) (d$i))}"
   unfolding expand_set_eq and Int_iff and mem_interval
-  by (auto simp add: Cart_lambda_beta' Arith_Tools.less_divide_eq_number_of1 intro!: bexI)
+  by (auto simp add: less_divide_eq_number_of1 intro!: bexI)
 
 (* Moved interval_open_subset_closed a bit upwards *)
 
@@ -4475,54 +4456,54 @@
  "a < x \<Longrightarrow> x < b ==> (\<exists>d>0. \<forall>x'. abs(x' - x) < d --> a < x' \<and> x' < b)"
   by(rule_tac x="min (x - a) (b - x)" in exI, auto)
 
-lemma open_interval: fixes a :: "real^'n" shows "open {a<..<b}"
+lemma open_interval: fixes a :: "real^'n::finite" shows "open {a<..<b}"
 proof-
   { fix x assume x:"x\<in>{a<..<b}"
-    { fix i assume "i\<in>dimset x"
-      hence "\<exists>d>0. \<forall>x'. abs (x' - (x$i)) < d \<longrightarrow> a$i < x' \<and> x' < b$i"
-	using x[unfolded mem_interval, THEN bspec[where x=i]]
+    { fix i
+      have "\<exists>d>0. \<forall>x'. abs (x' - (x$i)) < d \<longrightarrow> a$i < x' \<and> x' < b$i"
+	using x[unfolded mem_interval, THEN spec[where x=i]]
 	using open_interval_lemma[of "a$i" "x$i" "b$i"] by auto  }
 
-    hence "\<forall>i\<in>dimset x. \<exists>d>0. \<forall>x'. abs (x' - (x$i)) < d \<longrightarrow> a$i < x' \<and> x' < b$i" by auto
-    then obtain d where d:"\<forall>i\<in>dimset x. 0 < d i \<and> (\<forall>x'. \<bar>x' - x $ i\<bar> < d i \<longrightarrow> a $ i < x' \<and> x' < b $ i)"
-      using bchoice[of "dimset x" "\<lambda>i d. d>0 \<and> (\<forall>x'. \<bar>x' - x $ i\<bar> < d \<longrightarrow> a $ i < x' \<and> x' < b $ i)"] by auto
-
-    let ?d = "Min (d ` dimset x)"
-    have **:"finite (d ` dimset x)" "d ` dimset x \<noteq> {}" using dimindex_ge_1[of "UNIV::'n set"] by auto
+    hence "\<forall>i. \<exists>d>0. \<forall>x'. abs (x' - (x$i)) < d \<longrightarrow> a$i < x' \<and> x' < b$i" by auto
+    then obtain d where d:"\<forall>i. 0 < d i \<and> (\<forall>x'. \<bar>x' - x $ i\<bar> < d i \<longrightarrow> a $ i < x' \<and> x' < b $ i)"
+      using bchoice[of "UNIV" "\<lambda>i d. d>0 \<and> (\<forall>x'. \<bar>x' - x $ i\<bar> < d \<longrightarrow> a $ i < x' \<and> x' < b $ i)"] by auto
+
+    let ?d = "Min (range d)"
+    have **:"finite (range d)" "range d \<noteq> {}" by auto
     have "?d>0" unfolding Min_gr_iff[OF **] using d by auto
     moreover
     { fix x' assume as:"dist x' x < ?d"
-      { fix i assume i:"i \<in> dimset x"
+      { fix i
 	have "\<bar>x'$i - x $ i\<bar> < d i"
-	  using norm_bound_component_lt[OF as[unfolded dist_def], THEN bspec[where x=i], OF i]
-	  unfolding vector_minus_component[OF i] and Min_gr_iff[OF **] using i by auto
-	hence "a $ i < x' $ i" "x' $ i < b $ i" using d[THEN bspec[where x=i], OF i] by auto  }
+	  using norm_bound_component_lt[OF as[unfolded dist_def], of i]
+	  unfolding vector_minus_component and Min_gr_iff[OF **] by auto
+	hence "a $ i < x' $ i" "x' $ i < b $ i" using d[THEN spec[where x=i]] by auto  }
       hence "a < x' \<and> x' < b" unfolding vector_less_def by auto  }
-    ultimately have "\<exists>e>0. \<forall>x'. dist x' x < e \<longrightarrow> x' \<in> {a<..<b}" by auto
+    ultimately have "\<exists>e>0. \<forall>x'. dist x' x < e \<longrightarrow> x' \<in> {a<..<b}" by (auto, rule_tac x="?d" in exI, simp)
   }
   thus ?thesis unfolding open_def using open_interval_lemma by auto
 qed
 
-lemma closed_interval: fixes a :: "real^'n" shows "closed {a .. b}"
+lemma closed_interval: fixes a :: "real^'n::finite" shows "closed {a .. b}"
 proof-
-  { fix x i assume i:"i\<in>dimset x" and as:"\<forall>e>0. \<exists>x'\<in>{a..b}. x' \<noteq> x \<and> dist x' x < e"(* and xab:"a$i > x$i \<or> b$i < x$i"*)
+  { fix x i assume as:"\<forall>e>0. \<exists>x'\<in>{a..b}. x' \<noteq> x \<and> dist x' x < e"(* and xab:"a$i > x$i \<or> b$i < x$i"*)
     { assume xa:"a$i > x$i"
       with as obtain y where y:"y\<in>{a..b}" "y \<noteq> x" "dist y x < a$i - x$i" by(erule_tac x="a$i - x$i" in allE)auto
       hence False unfolding mem_interval and dist_def
-	using component_le_norm[OF i, of "y-x", unfolded vector_minus_component[OF i]] and i and xa by(auto elim!: ballE[where x=i])
+	using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xa by(auto elim!: allE[where x=i])
     } hence "a$i \<le> x$i" by(rule ccontr)auto
     moreover
     { assume xb:"b$i < x$i"
       with as obtain y where y:"y\<in>{a..b}" "y \<noteq> x" "dist y x < x$i - b$i" by(erule_tac x="x$i - b$i" in allE)auto
       hence False unfolding mem_interval and dist_def
-	using component_le_norm[OF i, of "y-x", unfolded vector_minus_component[OF i]] and i and xb by(auto elim!: ballE[where x=i])
+	using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xb by(auto elim!: allE[where x=i])
     } hence "x$i \<le> b$i" by(rule ccontr)auto
     ultimately
     have "a $ i \<le> x $ i \<and> x $ i \<le> b $ i" by auto }
   thus ?thesis unfolding closed_limpt islimpt_approachable mem_interval by auto
 qed
 
-lemma interior_closed_interval: fixes a :: "real^'n" shows
+lemma interior_closed_interval: fixes a :: "real^'n::finite" shows
  "interior {a .. b} = {a<..<b}" (is "?L = ?R")
 proof(rule subset_antisym)
   show "?R \<subseteq> ?L" using interior_maximal[OF interval_open_subset_closed open_interval] by auto
@@ -4530,85 +4511,87 @@
   { fix x assume "\<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> {a..b}"
     then obtain s where s:"open s" "x \<in> s" "s \<subseteq> {a..b}" by auto
     then obtain e where "e>0" and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> {a..b}" unfolding open_def and subset_eq by auto
-    { fix i assume i:"i\<in>dimset x"
+    { fix i
       have "dist (x - (e / 2) *s basis i) x < e"
 	   "dist (x + (e / 2) *s basis i) x < e"
 	unfolding dist_def apply auto
-	unfolding norm_minus_cancel and norm_mul using norm_basis[OF i] and `e>0` by auto
+	unfolding norm_minus_cancel and norm_mul using norm_basis[of i] and `e>0` by auto
       hence "a $ i \<le> (x - (e / 2) *s basis i) $ i"
                     "(x + (e / 2) *s basis i) $ i \<le> b $ i"
 	using e[THEN spec[where x="x - (e/2) *s basis i"]]
 	and   e[THEN spec[where x="x + (e/2) *s basis i"]]
-	unfolding mem_interval using i by auto
+	unfolding mem_interval by (auto elim!: allE[where x=i])
       hence "a $ i < x $ i" and "x $ i < b $ i"
-	unfolding vector_minus_component[OF i] and vector_add_component[OF i]
-	unfolding vector_smult_component[OF i] and basis_component[OF i] using `e>0` by auto   }
+	unfolding vector_minus_component and vector_add_component
+	unfolding vector_smult_component and basis_component using `e>0` by auto   }
     hence "x \<in> {a<..<b}" unfolding mem_interval by auto  }
   thus "?L \<subseteq> ?R" unfolding interior_def and subset_eq by auto
 qed
 
-lemma bounded_closed_interval: fixes a :: "real^'n" shows
+lemma bounded_closed_interval: fixes a :: "real^'n::finite" shows
  "bounded {a .. b}"
 proof-
-  let ?b = "\<Sum>i\<in>dimset a. \<bar>a$i\<bar> + \<bar>b$i\<bar>"
-  { fix x::"real^'n" assume x:"\<forall>i\<in>dimset a. a $ i \<le> x $ i \<and> x $ i \<le> b $ i"
-    { fix i assume "i\<in>dimset a"
-      hence "\<bar>x$i\<bar> \<le> \<bar>a$i\<bar> + \<bar>b$i\<bar>" using x[THEN bspec[where x=i]] by auto  }
-    hence "(\<Sum>i\<in>dimset a. \<bar>x $ i\<bar>) \<le> ?b" by(rule setsum_mono)auto
+  let ?b = "\<Sum>i\<in>UNIV. \<bar>a$i\<bar> + \<bar>b$i\<bar>"
+  { fix x::"real^'n" assume x:"\<forall>i. a $ i \<le> x $ i \<and> x $ i \<le> b $ i"
+    { fix i
+      have "\<bar>x$i\<bar> \<le> \<bar>a$i\<bar> + \<bar>b$i\<bar>" using x[THEN spec[where x=i]] by auto  }
+    hence "(\<Sum>i\<in>UNIV. \<bar>x $ i\<bar>) \<le> ?b" by(rule setsum_mono)
     hence "norm x \<le> ?b" using norm_le_l1[of x] by auto  }
   thus ?thesis unfolding interval and bounded_def by auto
 qed
 
-lemma bounded_interval: fixes a :: "real^'n" shows
+lemma bounded_interval: fixes a :: "real^'n::finite" shows
  "bounded {a .. b} \<and> bounded {a<..<b}"
   using bounded_closed_interval[of a b]
   using interval_open_subset_closed[of a b]
   using bounded_subset[of "{a..b}" "{a<..<b}"]
   by simp
 
-lemma not_interval_univ: fixes a :: "real^'n" shows
+lemma not_interval_univ: fixes a :: "real^'n::finite" shows
  "({a .. b} \<noteq> UNIV) \<and> ({a<..<b} \<noteq> UNIV)"
   using bounded_interval[of a b]
   by auto
 
-lemma compact_interval: fixes a :: "real^'n" shows
+lemma compact_interval: fixes a :: "real^'n::finite" shows
  "compact {a .. b}"
   using bounded_closed_imp_compact using bounded_interval[of a b] using closed_interval[of a b] by auto
 
-lemma open_interval_midpoint: fixes a :: "real^'n"
+lemma open_interval_midpoint: fixes a :: "real^'n::finite"
   assumes "{a<..<b} \<noteq> {}" shows "((1/2) *s (a + b)) \<in> {a<..<b}"
 proof-
-  { fix i assume i:"i\<in>dimset a"
-    hence "a $ i < ((1 / 2) *s (a + b)) $ i \<and> ((1 / 2) *s (a + b)) $ i < b $ i"
-      using assms[unfolded interval_ne_empty, THEN bspec[where x=i]]
-      unfolding vector_smult_component[OF i] and vector_add_component[OF i]
-      by(auto simp add: Arith_Tools.less_divide_eq_number_of1)  }
+  { fix i
+    have "a $ i < ((1 / 2) *s (a + b)) $ i \<and> ((1 / 2) *s (a + b)) $ i < b $ i"
+      using assms[unfolded interval_ne_empty, THEN spec[where x=i]]
+      unfolding vector_smult_component and vector_add_component
+      by(auto simp add: less_divide_eq_number_of1)  }
   thus ?thesis unfolding mem_interval by auto
 qed
 
-lemma open_closed_interval_convex: fixes x :: "real^'n"
+lemma open_closed_interval_convex: fixes x :: "real^'n::finite"
   assumes x:"x \<in> {a<..<b}" and y:"y \<in> {a .. b}" and e:"0 < e" "e \<le> 1"
   shows "(e *s x + (1 - e) *s y) \<in> {a<..<b}"
 proof-
-  { fix i assume i:"i\<in>dimset a"
+  { fix i
     have "a $ i = e * a$i + (1 - e) * a$i" unfolding left_diff_distrib by simp
     also have "\<dots> < e * x $ i + (1 - e) * y $ i" apply(rule add_less_le_mono)
       using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all
-      using x i unfolding mem_interval  apply(erule_tac x=i in ballE) apply simp_all
-      using y i unfolding mem_interval  apply(erule_tac x=i in ballE) by simp_all
-    finally have "a $ i < (e *s x + (1 - e) *s y) $ i" using i by (auto simp add: vector_add_component vector_smult_component)
+      using x unfolding mem_interval  apply simp
+      using y unfolding mem_interval  apply simp
+      done
+    finally have "a $ i < (e *s x + (1 - e) *s y) $ i" by auto
     moreover {
     have "b $ i = e * b$i + (1 - e) * b$i" unfolding left_diff_distrib by simp
     also have "\<dots> > e * x $ i + (1 - e) * y $ i" apply(rule add_less_le_mono)
       using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all
-      using x i unfolding mem_interval  apply(erule_tac x=i in ballE) apply simp_all
-      using y i unfolding mem_interval  apply(erule_tac x=i in ballE) by simp_all
-    finally have "(e *s x + (1 - e) *s y) $ i < b $ i" using i by (auto simp add: vector_add_component vector_smult_component)
+      using x unfolding mem_interval  apply simp
+      using y unfolding mem_interval  apply simp
+      done
+    finally have "(e *s x + (1 - e) *s y) $ i < b $ i" by auto
     } ultimately have "a $ i < (e *s x + (1 - e) *s y) $ i \<and> (e *s x + (1 - e) *s y) $ i < b $ i" by auto }
   thus ?thesis unfolding mem_interval by auto
 qed
 
-lemma closure_open_interval: fixes a :: "real^'n"
+lemma closure_open_interval: fixes a :: "real^'n::finite"
   assumes "{a<..<b} \<noteq> {}"
   shows "closure {a<..<b} = {a .. b}"
 proof-
@@ -4639,15 +4622,15 @@
   thus ?thesis using closure_minimal[OF interval_open_subset_closed closed_interval, of a b] by blast
 qed
 
-lemma bounded_subset_open_interval_symmetric: fixes s::"(real^'n) set"
+lemma bounded_subset_open_interval_symmetric: fixes s::"(real^'n::finite) set"
   assumes "bounded s"  shows "\<exists>a. s \<subseteq> {-a<..<a}"
 proof-
   obtain b where "b>0" and b:"\<forall>x\<in>s. norm x \<le> b" using assms[unfolded bounded_pos] by auto
   def a \<equiv> "(\<chi> i. b+1)::real^'n"
   { fix x assume "x\<in>s"
-    fix i assume i:"i\<in>dimset a"
-    have "(-a)$i < x$i" and "x$i < a$i" using b[THEN bspec[where x=x], OF `x\<in>s`] and component_le_norm[OF i, of x]
-      unfolding vector_uminus_component[OF i] and a_def and Cart_lambda_beta'[OF i] by auto
+    fix i
+    have "(-a)$i < x$i" and "x$i < a$i" using b[THEN bspec[where x=x], OF `x\<in>s`] and component_le_norm[of x i]
+      unfolding vector_uminus_component and a_def and Cart_lambda_beta by auto
   }
   thus ?thesis by(auto intro: exI[where x=a] simp add: vector_less_def)
 qed
@@ -4679,30 +4662,32 @@
   case False thus ?thesis unfolding frontier_def and closure_open_interval[OF False] and interior_open[OF open_interval] by auto
 qed
 
-lemma inter_interval_mixed_eq_empty: fixes a :: "real^'n"
+lemma inter_interval_mixed_eq_empty: fixes a :: "real^'n::finite"
   assumes "{c<..<d} \<noteq> {}"  shows "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> {a<..<b} \<inter> {c<..<d} = {}"
   unfolding closure_open_interval[OF assms, THEN sym] unfolding open_inter_closure_eq_empty[OF open_interval] ..
 
 
 (* Some special cases for intervals in R^1.                                  *)
 
-lemma dim1: "dimindex (UNIV::(1 set)) = 1"
-unfolding dimindex_def
-by simp
+lemma all_1: "(\<forall>x::1. P x) \<longleftrightarrow> P 1"
+  by (metis num1_eq_iff)
+
+lemma ex_1: "(\<exists>x::1. P x) \<longleftrightarrow> P 1"
+  by auto (metis num1_eq_iff)
 
 lemma interval_cases_1: fixes x :: "real^1" shows
  "x \<in> {a .. b} ==> x \<in> {a<..<b} \<or> (x = a) \<or> (x = b)"
-  by(simp add:  Cart_eq vector_less_def vector_less_eq_def dim1, auto)
+  by(simp add:  Cart_eq vector_less_def vector_less_eq_def all_1, auto)
 
 lemma in_interval_1: fixes x :: "real^1" shows
  "(x \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b) \<and>
   (x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> dest_vec1 x < dest_vec1 b)"
-by(simp add: Cart_eq vector_less_def vector_less_eq_def dim1 dest_vec1_def)
+by(simp add: Cart_eq vector_less_def vector_less_eq_def all_1 dest_vec1_def)
 
 lemma interval_eq_empty_1: fixes a :: "real^1" shows
   "{a .. b} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a"
   "{a<..<b} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a"
-  unfolding interval_eq_empty and dim1 and dest_vec1_def by auto
+  unfolding interval_eq_empty and ex_1 and dest_vec1_def by auto
 
 lemma subset_interval_1: fixes a :: "real^1" shows
  "({a .. b} \<subseteq> {c .. d} \<longleftrightarrow>  dest_vec1 b < dest_vec1 a \<or>
@@ -4713,7 +4698,7 @@
                 dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a < dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)"
  "({a<..<b} \<subseteq> {c<..<d} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or>
                 dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a < dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)"
-  unfolding subset_interval[of a b c d] unfolding forall_dimindex_1 and dest_vec1_def by auto
+  unfolding subset_interval[of a b c d] unfolding all_1 and dest_vec1_def by auto
 
 lemma eq_interval_1: fixes a :: "real^1" shows
  "{a .. b} = {c .. d} \<longleftrightarrow>
@@ -4729,37 +4714,37 @@
   "{a .. b} \<inter> {c<..<d} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a \<or> dest_vec1 d \<le> dest_vec1 c  \<or>  dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
   "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or> dest_vec1 d < dest_vec1 c  \<or>  dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
   "{a<..<b} \<inter> {c<..<d} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or> dest_vec1 d \<le> dest_vec1 c  \<or>  dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
-  unfolding disjoint_interval and dest_vec1_def and dim1 by auto
+  unfolding disjoint_interval and dest_vec1_def ex_1 by auto
 
 lemma open_closed_interval_1: fixes a :: "real^1" shows
  "{a<..<b} = {a .. b} - {a, b}"
-  unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and dim1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto
+  unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and all_1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto
 
 lemma closed_open_interval_1: "dest_vec1 (a::real^1) \<le> dest_vec1 b ==> {a .. b} = {a<..<b} \<union> {a,b}"
-  unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and dim1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto
+  unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and all_1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto
 
 (* Some stuff for half-infinite intervals too; FIXME: notation?  *)
 
-lemma closed_interval_left: fixes b::"real^'n"
-  shows "closed {x::real^'n. \<forall>i \<in> dimset x. x$i \<le> b$i}"
+lemma closed_interval_left: fixes b::"real^'n::finite"
+  shows "closed {x::real^'n. \<forall>i. x$i \<le> b$i}"
 proof-
-  { fix i assume i:"i\<in>dimset b"
-    fix x::"real^'n" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i\<in>dimset b. x $ i \<le> b $ i}. x' \<noteq> x \<and> dist x' x < e"
+  { fix i
+    fix x::"real^'n" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i. x $ i \<le> b $ i}. x' \<noteq> x \<and> dist x' x < e"
     { assume "x$i > b$i"
-      then obtain y where "y $ i \<le> b $ i"  "y \<noteq> x"  "dist y x < x$i - b$i" using x[THEN spec[where x="x$i - b$i"]] and i by (auto, erule_tac x=i in ballE)auto
-      hence False using component_le_norm[OF i, of "y - x"] unfolding dist_def and vector_minus_component[OF i] by auto   }
+      then obtain y where "y $ i \<le> b $ i"  "y \<noteq> x"  "dist y x < x$i - b$i" using x[THEN spec[where x="x$i - b$i"]] by auto
+      hence False using component_le_norm[of "y - x" i] unfolding dist_def and vector_minus_component by auto   }
     hence "x$i \<le> b$i" by(rule ccontr)auto  }
   thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
 qed
 
-lemma closed_interval_right: fixes a::"real^'n"
-  shows "closed {x::real^'n. \<forall>i \<in> dimset x. a$i \<le> x$i}"
+lemma closed_interval_right: fixes a::"real^'n::finite"
+  shows "closed {x::real^'n. \<forall>i. a$i \<le> x$i}"
 proof-
-  { fix i assume i:"i\<in>dimset a"
-    fix x::"real^'n" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i\<in>dimset a. a $ i \<le> x $ i}. x' \<noteq> x \<and> dist x' x < e"
+  { fix i
+    fix x::"real^'n" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i. a $ i \<le> x $ i}. x' \<noteq> x \<and> dist x' x < e"
     { assume "a$i > x$i"
-      then obtain y where "a $ i \<le> y $ i"  "y \<noteq> x"  "dist y x < a$i - x$i" using x[THEN spec[where x="a$i - x$i"]] and i by(auto, erule_tac x=i in ballE)auto
-      hence False using component_le_norm[OF i, of "y - x"] unfolding dist_def and vector_minus_component[OF i] by auto   }
+      then obtain y where "a $ i \<le> y $ i"  "y \<noteq> x"  "dist y x < a$i - x$i" using x[THEN spec[where x="a$i - x$i"]] by auto
+      hence False using component_le_norm[of "y - x" i] unfolding dist_def and vector_minus_component by auto   }
     hence "a$i \<le> x$i" by(rule ccontr)auto  }
   thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
 qed
@@ -4768,13 +4753,13 @@
 
 definition "is_interval s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> x \<in> s)"
 
-lemma is_interval_interval: fixes a::"real^'n" shows
+lemma is_interval_interval: fixes a::"real^'n::finite" shows
   "is_interval {a<..<b}" "is_interval {a .. b}"
   unfolding is_interval_def apply(auto simp add: vector_less_def vector_less_eq_def)
-  apply(erule_tac x=i in ballE)+ apply simp+
-  apply(erule_tac x=i in ballE)+ apply simp+
-  apply(erule_tac x=i in ballE)+ apply simp+
-  apply(erule_tac x=i in ballE)+ apply simp+
+  apply(erule_tac x=i in allE)+ apply simp
+  apply(erule_tac x=i in allE)+ apply simp
+  apply(erule_tac x=i in allE)+ apply simp
+  apply(erule_tac x=i in allE)+ apply simp
   done
 
 lemma is_interval_empty:
@@ -4789,7 +4774,7 @@
 
 subsection{* Closure of halfspaces and hyperplanes.                                    *}
 
-lemma Lim_vec1_dot: fixes f :: "real^'m \<Rightarrow> real^'n"
+lemma Lim_vec1_dot: fixes f :: "real^'m \<Rightarrow> real^'n::finite"
   assumes "(f ---> l) net"  shows "((vec1 o (\<lambda>y. a \<bullet> (f y))) ---> vec1(a \<bullet> l)) net"
 proof(cases "a = vec 0")
   case True thus ?thesis using dot_lzero and Lim_const[of 0 net] unfolding vec1_vec and o_def by auto
@@ -4821,14 +4806,14 @@
   using continuous_at_vec1_dot
   by auto
 
-lemma closed_halfspace_le: fixes a::"real^'n"
+lemma closed_halfspace_le: fixes a::"real^'n::finite"
   shows "closed {x. a \<bullet> x \<le> b}"
 proof-
   have *:"{x \<in> UNIV. (vec1 \<circ> op \<bullet> a) x \<in> vec1 ` {r. \<exists>x. a \<bullet> x = r \<and> r \<le> b}} = {x. a \<bullet> x \<le> b}" by auto
-  let ?T = "{x::real^1. (\<forall>i\<in>dimset x. x$i \<le> (vec1 b)$i)}"
+  let ?T = "{x::real^1. (\<forall>i. x$i \<le> (vec1 b)$i)}"
   have "closed ?T" using closed_interval_left[of "vec1 b"] by simp
-  moreover have "vec1 ` {r. \<exists>x. a \<bullet> x = r \<and> r \<le> b} = range (vec1 \<circ> op \<bullet> a) \<inter> ?T" unfolding dim1
-    unfolding image_def apply auto unfolding vec1_component[unfolded One_nat_def] by auto
+  moreover have "vec1 ` {r. \<exists>x. a \<bullet> x = r \<and> r \<le> b} = range (vec1 \<circ> op \<bullet> a) \<inter> ?T" unfolding all_1
+    unfolding image_def by auto
   ultimately have "\<exists>T. closed T \<and> vec1 ` {r. \<exists>x. a \<bullet> x = r \<and> r \<le> b} = range (vec1 \<circ> op \<bullet> a) \<inter> T" by auto
   hence "closedin euclidean {x \<in> UNIV. (vec1 \<circ> op \<bullet> a) x \<in> vec1 ` {r. \<exists>x. a \<bullet> x = r \<and> r \<le> b}}"
     using continuous_on_vec1_dot[of UNIV a, unfolded continuous_on_closed subtopology_UNIV] unfolding closedin_closed
@@ -4846,11 +4831,11 @@
 qed
 
 lemma closed_halfspace_component_le:
-  assumes "i \<in> {1 .. dimindex (UNIV::'n set)}" shows "closed {x::real^'n. x$i \<le> a}"
+  shows "closed {x::real^'n::finite. x$i \<le> a}"
   using closed_halfspace_le[of "(basis i)::real^'n" a] unfolding dot_basis[OF assms] by auto
 
 lemma closed_halfspace_component_ge:
-  assumes "i \<in> {1 .. dimindex (UNIV::'n set)}" shows "closed {x::real^'n. x$i \<ge> a}"
+  shows "closed {x::real^'n::finite. x$i \<ge> a}"
   using closed_halfspace_ge[of a "(basis i)::real^'n"] unfolding dot_basis[OF assms] by auto
 
 text{* Openness of halfspaces.                                                   *}
@@ -4868,48 +4853,45 @@
 qed
 
 lemma open_halfspace_component_lt:
-  assumes "i \<in> {1 .. dimindex(UNIV::'n set)}" shows "open {x::real^'n. x$i < a}"
+  shows "open {x::real^'n::finite. x$i < a}"
   using open_halfspace_lt[of "(basis i)::real^'n" a] unfolding dot_basis[OF assms] by auto
 
 lemma open_halfspace_component_gt:
-  assumes "i \<in> {1 .. dimindex(UNIV::'n set)}" shows "open {x::real^'n. x$i  > a}"
+  shows "open {x::real^'n::finite. x$i  > a}"
   using open_halfspace_gt[of a "(basis i)::real^'n"] unfolding dot_basis[OF assms] by auto
 
 text{* This gives a simple derivation of limit component bounds.                 *}
 
-lemma Lim_component_le: fixes f :: "'a \<Rightarrow> real^'n"
+lemma Lim_component_le: fixes f :: "'a \<Rightarrow> real^'n::finite"
   assumes "(f ---> l) net" "\<not> (trivial_limit net)"  "eventually (\<lambda>x. f(x)$i \<le> b) net"
-  and i:"i\<in> {1 .. dimindex(UNIV::'n set)}"
   shows "l$i \<le> b"
 proof-
-  { fix x have "x \<in> {x::real^'n. basis i \<bullet> x \<le> b} \<longleftrightarrow> x$i \<le> b" unfolding dot_basis[OF i] by auto } note * = this
+  { fix x have "x \<in> {x::real^'n. basis i \<bullet> x \<le> b} \<longleftrightarrow> x$i \<le> b" unfolding dot_basis by auto } note * = this
   show ?thesis using Lim_in_closed_set[of "{x. basis i \<bullet> x \<le> b}" f net l] unfolding *
     using closed_halfspace_le[of "(basis i)::real^'n" b] and assms(1,2,3) by auto
 qed
 
-lemma Lim_component_ge: fixes f :: "'a \<Rightarrow> real^'n"
+lemma Lim_component_ge: fixes f :: "'a \<Rightarrow> real^'n::finite"
   assumes "(f ---> l) net"  "\<not> (trivial_limit net)"  "eventually (\<lambda>x. b \<le> (f x)$i) net"
-  and i:"i\<in> {1 .. dimindex(UNIV::'n set)}"
   shows "b \<le> l$i"
 proof-
-  { fix x have "x \<in> {x::real^'n. basis i \<bullet> x \<ge> b} \<longleftrightarrow> x$i \<ge> b" unfolding dot_basis[OF i] by auto } note * = this
+  { fix x have "x \<in> {x::real^'n. basis i \<bullet> x \<ge> b} \<longleftrightarrow> x$i \<ge> b" unfolding dot_basis by auto } note * = this
   show ?thesis using Lim_in_closed_set[of "{x. basis i \<bullet> x \<ge> b}" f net l] unfolding *
     using closed_halfspace_ge[of b "(basis i)::real^'n"] and assms(1,2,3) by auto
 qed
 
-lemma Lim_component_eq: fixes f :: "'a \<Rightarrow> real^'n"
+lemma Lim_component_eq: fixes f :: "'a \<Rightarrow> real^'n::finite"
   assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\<lambda>x. f(x)$i = b) net"
-  and i:"i\<in> {1 .. dimindex(UNIV::'n set)}"
   shows "l$i = b"
-  using ev[unfolded order_eq_iff eventually_and] using Lim_component_ge[OF net, of b i] and Lim_component_le[OF net, of i b] using i by auto
+  using ev[unfolded order_eq_iff eventually_and] using Lim_component_ge[OF net, of b i] and Lim_component_le[OF net, of i b] by auto
 
 lemma Lim_drop_le: fixes f :: "'a \<Rightarrow> real^1" shows
   "(f ---> l) net \<Longrightarrow> ~(trivial_limit net) \<Longrightarrow> eventually (\<lambda>x. dest_vec1 (f x) \<le> b) net ==> dest_vec1 l \<le> b"
-  using Lim_component_le[of f l net 1 b] unfolding dest_vec1_def and dim1 by auto
+  using Lim_component_le[of f l net 1 b] unfolding dest_vec1_def by auto
 
 lemma Lim_drop_ge: fixes f :: "'a \<Rightarrow> real^1" shows
  "(f ---> l) net \<Longrightarrow> ~(trivial_limit net) \<Longrightarrow> eventually (\<lambda>x. b \<le> dest_vec1 (f x)) net ==> b \<le> dest_vec1 l"
-  using Lim_component_ge[of f l net b 1] unfolding dest_vec1_def and dim1 by auto
+  using Lim_component_ge[of f l net b 1] unfolding dest_vec1_def by auto
 
 text{* Limits relative to a union.                                               *}
 
@@ -4926,7 +4908,7 @@
   using assms unfolding continuous_on unfolding Lim_within_union
   unfolding Lim unfolding trivial_limit_within unfolding closed_limpt by auto
 
-lemma continuous_on_cases: fixes g :: "real^'m \<Rightarrow> real ^'n"
+lemma continuous_on_cases: fixes g :: "real^'m::finite \<Rightarrow> real ^'n::finite"
   assumes "closed s" "closed t" "continuous_on s f" "continuous_on t g"
           "\<forall>x. (x\<in>s \<and> \<not> P x) \<or> (x \<in> t \<and> P x) \<longrightarrow> f x = g x"
   shows "continuous_on (s \<union> t) (\<lambda>x. if P x then f x else g x)"
@@ -4943,7 +4925,7 @@
 
 text{* Some more convenient intermediate-value theorem formulations.             *}
 
-lemma connected_ivt_hyperplane: fixes y :: "real^'n"
+lemma connected_ivt_hyperplane: fixes y :: "real^'n::finite"
   assumes "connected s" "x \<in> s" "y \<in> s" "a \<bullet> x \<le> b" "b \<le> a \<bullet> y"
   shows "\<exists>z \<in> s. a \<bullet> z = b"
 proof(rule ccontr)
@@ -4956,8 +4938,8 @@
   ultimately show False using assms(1)[unfolded connected_def not_ex, THEN spec[where x="?A"], THEN spec[where x="?B"]] and assms(2-5) by auto
 qed
 
-lemma connected_ivt_component: fixes x::"real^'n" shows
- "connected s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> k \<in> dimset x \<Longrightarrow> x$k \<le> a \<Longrightarrow> a \<le> y$k \<Longrightarrow> (\<exists>z\<in>s.  z$k = a)"
+lemma connected_ivt_component: fixes x::"real^'n::finite" shows
+ "connected s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x$k \<le> a \<Longrightarrow> a \<le> y$k \<Longrightarrow> (\<exists>z\<in>s.  z$k = a)"
   using connected_ivt_hyperplane[of s x y "(basis k)::real^'n" a] by (auto simp add: dot_basis)
 
 text{* Also more convenient formulations of monotone convergence.                *}
@@ -4982,7 +4964,7 @@
      (\<forall>x\<in>s. (g(f x) = x)) \<and> (f ` s = t) \<and> continuous_on s f \<and>
      (\<forall>y\<in>t. (f(g y) = y)) \<and> (g ` t = s) \<and> continuous_on t g"
 
-definition homeomorphic :: "((real^'a) set) \<Rightarrow> ((real^'b) set) \<Rightarrow> bool" (infixr "homeomorphic" 60) where
+definition homeomorphic :: "((real^'a::finite) set) \<Rightarrow> ((real^'b::finite) set) \<Rightarrow> bool" (infixr "homeomorphic" 60) where
   homeomorphic_def: "s homeomorphic t \<equiv> (\<exists>f g. homeomorphism s t f g)"
 
 lemma homeomorphic_refl: "s homeomorphic s"
@@ -5103,7 +5085,7 @@
     using homeomorphic_translation[of "(\<lambda>x. c *s x) ` s" a] unfolding * by auto
 qed
 
-lemma homeomorphic_balls: fixes a b ::"real^'a"
+lemma homeomorphic_balls: fixes a b ::"real^'a::finite"
   assumes "0 < d"  "0 < e"
   shows "(ball a d) homeomorphic  (ball b e)" (is ?th)
         "(cball a d) homeomorphic (cball b e)" (is ?cth)
@@ -5173,7 +5155,7 @@
 
 lemma dist_0_norm:"dist 0 x = norm x" unfolding dist_def by(auto simp add: norm_minus_cancel)
 
-lemma injective_imp_isometric: fixes f::"real^'m \<Rightarrow> real^'n"
+lemma injective_imp_isometric: fixes f::"real^'m::finite \<Rightarrow> real^'n::finite"
   assumes s:"closed s"  "subspace s"  and f:"linear f" "\<forall>x\<in>s. (f x = 0) \<longrightarrow> (x = 0)"
   shows "\<exists>e>0. \<forall>x\<in>s. norm (f x) \<ge> e * norm(x)"
 proof(cases "s \<subseteq> {0::real^'m}")
@@ -5235,24 +5217,23 @@
 subsection{* Some properties of a canonical subspace.                                  *}
 
 lemma subspace_substandard:
- "subspace {x::real^'n. (\<forall>i \<in> dimset x. d < i \<longrightarrow> x$i = 0)}"
+ "subspace {x::real^'n. (\<forall>i. P i \<longrightarrow> x$i = 0)}"
   unfolding subspace_def by(auto simp add: vector_add_component vector_smult_component elim!: ballE)
 
 lemma closed_substandard:
- "closed {x::real^'n. \<forall>i \<in> dimset x. d < i --> x$i = 0}" (is "closed ?A")
+ "closed {x::real^'n::finite. \<forall>i. P i --> x$i = 0}" (is "closed ?A")
 proof-
-  let ?D = "{Suc d..dimindex(UNIV::('n set))}"
+  let ?D = "{i. P i}"
   let ?Bs = "{{x::real^'n. basis i \<bullet> x = 0}| i. i \<in> ?D}"
   { fix x
     { assume "x\<in>?A"
-      hence x:"\<forall>i\<in>?D. d < i \<longrightarrow> x $ i = 0" by auto
+      hence x:"\<forall>i\<in>?D. x $ i = 0" by auto
       hence "x\<in> \<Inter> ?Bs" by(auto simp add: dot_basis x) }
     moreover
     { assume x:"x\<in>\<Inter>?Bs"
-      { fix i assume i:"i\<in>dimset x" and "d < i"
-	hence "i \<in> ?D" by auto
+      { fix i assume i:"i \<in> ?D"
 	then obtain B where BB:"B \<in> ?Bs" and B:"B = {x::real^'n. basis i \<bullet> x = 0}" by auto
-	hence "x $ i = 0" unfolding B unfolding dot_basis[OF i] using x by auto  }
+	hence "x $ i = 0" unfolding B using x unfolding dot_basis by auto  }
       hence "x\<in>?A" by auto }
     ultimately have "x\<in>?A \<longleftrightarrow> x\<in> \<Inter>?Bs" by auto }
   hence "?A = \<Inter> ?Bs" by auto
@@ -5260,40 +5241,40 @@
 qed
 
 lemma dim_substandard:
-  assumes "d \<le> dimindex(UNIV::'n set)"
-  shows "dim {x::real^'n. \<forall>i \<in> dimset x. d < i --> x$i = 0} = d" (is "dim ?A = d")
+  shows "dim {x::real^'n::finite. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0} = card d" (is "dim ?A = _")
 proof-
-  let ?D = "{1..dimindex (UNIV::'n set)}"
-  let ?B = "(basis::nat\<Rightarrow>real^'n) ` {1..d}"
-
-    let ?bas = "basis::nat \<Rightarrow> real^'n"
-
-  have "?B \<subseteq> ?A" by (auto simp add: basis_component)
+  let ?D = "UNIV::'n set"
+  let ?B = "(basis::'n\<Rightarrow>real^'n) ` d"
+
+    let ?bas = "basis::'n \<Rightarrow> real^'n"
+
+  have "?B \<subseteq> ?A" by auto
 
   moreover
   { fix x::"real^'n" assume "x\<in>?A"
-    hence "x\<in> span ?B"
+    with finite[of d]
+    have "x\<in> span ?B"
     proof(induct d arbitrary: x)
-      case 0 hence "x=0" unfolding Cart_eq by auto
+      case empty hence "x=0" unfolding Cart_eq by auto
       thus ?case using subspace_0[OF subspace_span[of "{}"]] by auto
     next
-      case (Suc n)
-      hence *:"\<forall>i\<in>?D. Suc n < i \<longrightarrow> x $ i = 0" by auto
-      have **:"{1..n} \<subseteq> {1..Suc n}" by auto
-      def y \<equiv> "x - x$(Suc n) *s basis (Suc n)"
-      have y:"x = y + (x$Suc n) *s basis (Suc n)" unfolding y_def by auto
-      { fix i assume i:"i\<in>?D" and i':"n < i"
-	hence "y $ i = 0" unfolding y_def unfolding vector_minus_component[OF i]
-	  and vector_smult_component[OF i] and basis_component[OF i] using i'
-	  using *[THEN bspec[where x=i]] by auto }
-      hence "y \<in> span (basis ` {1..Suc n})" using Suc(1)[of y]
-	using span_mono[of "?bas ` {1..n}" "?bas ` {1..Suc n}"]
+      case (insert k F)
+      hence *:"\<forall>i. i \<notin> insert k F \<longrightarrow> x $ i = 0" by auto
+      have **:"F \<subseteq> insert k F" by auto
+      def y \<equiv> "x - x$k *s basis k"
+      have y:"x = y + (x$k) *s basis k" unfolding y_def by auto
+      { fix i assume i':"i \<notin> F"
+	hence "y $ i = 0" unfolding y_def unfolding vector_minus_component
+	  and vector_smult_component and basis_component
+	  using *[THEN spec[where x=i]] by auto }
+      hence "y \<in> span (basis ` (insert k F))" using insert(3)
+	using span_mono[of "?bas ` F" "?bas ` (insert k F)"]
 	using image_mono[OF **, of basis] by auto
       moreover
-      have "basis (Suc n) \<in> span (?bas ` {1..Suc n})" by(rule span_superset, auto)
-      hence "x$(Suc n) *s basis (Suc n) \<in> span (?bas ` {1..Suc n})" using span_mul by auto
+      have "basis k \<in> span (?bas ` (insert k F))" by(rule span_superset, auto)
+      hence "x$k *s basis k \<in> span (?bas ` (insert k F))" using span_mul by auto
       ultimately
-      have "y + x$(Suc n) *s basis (Suc n) \<in> span (?bas ` {1..Suc n})"
+      have "y + x$k *s basis k \<in> span (?bas ` (insert k F))"
 	using span_add by auto
       thus ?case using y by auto
     qed
@@ -5306,27 +5287,39 @@
   hence "independent ?B" using independent_mono[OF independent_stdbasis, of ?B] and assms by auto
 
   moreover
-  have "{1..d} \<subseteq> ?D" unfolding subset_eq using assms by auto
-  hence *:"inj_on (basis::nat\<Rightarrow>real^'n) {1..d}" using subset_inj_on[OF basis_inj, of "{1..d}"] using assms by auto
-  have "?B hassize d" unfolding hassize_def and card_image[OF *] by auto
-
-  ultimately show ?thesis using dim_unique[of "basis ` {1..d}" ?A] by auto
+  have "d \<subseteq> ?D" unfolding subset_eq using assms by auto
+  hence *:"inj_on (basis::'n\<Rightarrow>real^'n) d" using subset_inj_on[OF basis_inj, of "d"] by auto
+  have "?B hassize (card d)" unfolding hassize_def and card_image[OF *] by auto
+
+  ultimately show ?thesis using dim_unique[of "basis ` d" ?A] by auto
 qed
 
 text{* Hence closure and completeness of all subspaces.                          *}
 
-lemma closed_subspace: fixes s::"(real^'n) set"
+lemma closed_subspace_lemma: "n \<le> card (UNIV::'n::finite set) \<Longrightarrow> \<exists>A::'n set. card A = n"
+apply (induct n)
+apply (rule_tac x="{}" in exI, simp)
+apply clarsimp
+apply (subgoal_tac "\<exists>x. x \<notin> A")
+apply (erule exE)
+apply (rule_tac x="insert x A" in exI, simp)
+apply (subgoal_tac "A \<noteq> UNIV", auto)
+done
+
+lemma closed_subspace: fixes s::"(real^'n::finite) set"
   assumes "subspace s" shows "closed s"
 proof-
-  let ?t = "{x::real^'n. \<forall>i\<in>{1..dimindex (UNIV :: 'n set)}. dim s<i \<longrightarrow> x$i = 0}"
-  have "dim s \<le> dimindex (UNIV :: 'n set)" using dim_subset_univ by auto
+  have "dim s \<le> card (UNIV :: 'n set)" using dim_subset_univ by auto
+  then obtain d::"'n set" where t: "card d = dim s"
+    using closed_subspace_lemma by auto
+  let ?t = "{x::real^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0}"
   obtain f where f:"linear f"  "f ` ?t = s" "inj_on f ?t"
-    using subspace_isomorphism[OF subspace_substandard[of "dim s"] assms]
-    using dim_substandard[OF  dim_subset_univ[of s]] by auto
+    using subspace_isomorphism[OF subspace_substandard[of "\<lambda>i. i \<notin> d"] assms]
+    using dim_substandard[of d] and t by auto
   have "\<forall>x\<in>?t. f x = 0 \<longrightarrow> x = 0" using linear_0[OF f(1)] using f(3)[unfolded inj_on_def]
     by(erule_tac x=0 in ballE) auto
-  moreover have "closed ?t" using closed_substandard by auto
-  moreover have "subspace ?t" using subspace_substandard by auto
+  moreover have "closed ?t" using closed_substandard .
+  moreover have "subspace ?t" using subspace_substandard .
   ultimately show ?thesis using closed_injective_image_subspace[of ?t f]
     unfolding f(2) using f(1) by auto
 qed
@@ -5400,13 +5393,14 @@
   by metis
 
 lemma image_affinity_interval: fixes m::real
+  fixes a b c :: "real^'n::finite"
   shows "(\<lambda>x. m *s x + c) ` {a .. b} =
             (if {a .. b} = {} then {}
             else (if 0 \<le> m then {m *s a + c .. m *s b + c}
             else {m *s b + c .. m *s a + c}))"
 proof(cases "m=0")
   { fix x assume "x \<le> c" "c \<le> x"
-    hence "x=c" unfolding vector_less_eq_def and Cart_eq by(auto elim!: ballE)  }
+    hence "x=c" unfolding vector_less_eq_def and Cart_eq by (auto intro: order_antisym) }
   moreover case True
   moreover have "c \<in> {m *s a + c..m *s b + c}" unfolding True by(auto simp add: vector_less_eq_def)
   ultimately show ?thesis by auto
@@ -5425,14 +5419,14 @@
       unfolding image_iff Bex_def mem_interval vector_less_eq_def
       apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric]
 	intro!: exI[where x="(1 / m) *s (y - c)"])
-      by(auto elim!: ballE simp add: pos_le_divide_eq pos_divide_le_eq real_mult_commute)
+      by(auto simp add: pos_le_divide_eq pos_divide_le_eq real_mult_commute diff_le_iff)
   } moreover
   { fix y assume "m *s b + c \<le> y" "y \<le> m *s a + c" "m < 0"
     hence "y \<in> (\<lambda>x. m *s x + c) ` {a..b}"
       unfolding image_iff Bex_def mem_interval vector_less_eq_def
       apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric]
 	intro!: exI[where x="(1 / m) *s (y - c)"])
-      by(auto elim!: ballE simp add: neg_le_divide_eq neg_divide_le_eq real_mult_commute)
+      by(auto simp add: neg_le_divide_eq neg_divide_le_eq real_mult_commute diff_le_iff)
   }
   ultimately show ?thesis using False by auto
 qed
@@ -5569,7 +5563,7 @@
 lemma edelstein_fix:
   assumes s:"compact s" "s \<noteq> {}" and gs:"(g ` s) \<subseteq> s"
       and dist:"\<forall>x\<in>s. \<forall>y\<in>s. x \<noteq> y \<longrightarrow> dist (g x) (g y) < dist x y"
-  shows "\<exists>! x::real^'a\<in>s. g x = x"
+  shows "\<exists>! x::real^'a::finite\<in>s. g x = x"
 proof(cases "\<exists>x\<in>s. g x \<noteq> x")
   obtain x where "x\<in>s" using s(2) by auto
   case False hence g:"\<forall>x\<in>s. g x = x" by auto
@@ -5637,7 +5631,7 @@
     { assume as:"dist a b > dist (f n x) (f n y)"
       then obtain Na Nb where "\<forall>m\<ge>Na. dist (f (r m) x) a < (dist a b - dist (f n x) (f n y)) / 2"
 	and "\<forall>m\<ge>Nb. dist (f (r m) y) b < (dist a b - dist (f n x) (f n y)) / 2"
-	using lima limb unfolding h_def Lim_sequentially by (fastsimp simp del: Arith_Tools.less_divide_eq_number_of1)
+	using lima limb unfolding h_def Lim_sequentially by (fastsimp simp del: less_divide_eq_number_of1)
       hence "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) < dist a b - dist (f n x) (f n y)"
 	apply(erule_tac x="Na+Nb+n" in allE)
 	apply(erule_tac x="Na+Nb+n" in allE) apply simp
--- a/src/HOL/Library/Univ_Poly.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Univ_Poly.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Univariate Polynomials *}
 
 theory Univ_Poly
-imports Plain List
+imports Main
 begin
 
 text{* Application of polynomial as a function. *}
--- a/src/HOL/Library/While_Combinator.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/While_Combinator.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Library/While_Combinator.thy
-    ID:         $Id$
     Author:     Tobias Nipkow
     Copyright   2000 TU Muenchen
 *)
@@ -7,7 +6,7 @@
 header {* A general ``while'' combinator *}
 
 theory While_Combinator
-imports Plain "~~/src/HOL/Presburger"
+imports Main
 begin
 
 text {* 
--- a/src/HOL/Library/Zorn.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/Zorn.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -7,7 +7,7 @@
 header {* Zorn's Lemma *}
 
 theory Zorn
-imports "~~/src/HOL/Order_Relation"
+imports Order_Relation Main
 begin
 
 (* Define globally? In Set.thy? *)
--- a/src/HOL/Library/reflection.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Library/reflection.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -243,8 +243,8 @@
 		     insert (op =) (eq |> prop_of |> HOLogic.dest_Trueprop 
 			 |> HOLogic.dest_eq |> fst |> strip_comb 
 			 |> fst)) raw_eqs []
-  val tys = fold_rev (fn f => fn ts => (f |> fastype_of |> binder_types |> tl) 
-				    union ts) fs []
+  val tys = fold_rev (fn f => fold (insert (op =)) (f |> fastype_of |> binder_types |> tl) 
+				    ) fs []
   val _ = bds := AList.make (fn _ => ([],[])) tys
   val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt
   val thy = ProofContext.theory_of ctxt'
--- a/src/HOL/Lubs.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Lubs.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -6,7 +6,7 @@
 header{*Definitions of Upper Bounds and Least Upper Bounds*}
 
 theory Lubs
-imports Plain Main
+imports Main
 begin
 
 text{*Thanks to suggestions by James Margetson*}
--- a/src/HOL/Map.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Map.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Map.thy
-    ID:         $Id$
     Author:     Tobias Nipkow, based on a theory by David von Oheimb
     Copyright   1997-2003 TU Muenchen
 
--- a/src/HOL/Modelcheck/mucke_oracle.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Modelcheck/mucke_oracle.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -109,7 +109,7 @@
 (
 OldGoals.push_proof();
 OldGoals.goalw_cterm [] (cterm_of sign trm);
-by (SIMPSET' simp_tac 1);
+by (simp_tac (simpset_of sign) 1);
         let
         val if_tmp_result = result()
         in
--- a/src/HOL/NSA/StarDef.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/NSA/StarDef.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -23,7 +23,7 @@
 apply (rule nat_infinite)
 done
 
-interpretation FreeUltrafilterNat!: freeultrafilter FreeUltrafilterNat
+interpretation FreeUltrafilterNat: freeultrafilter FreeUltrafilterNat
 by (rule freeultrafilter_FreeUltrafilterNat)
 
 text {* This rule takes the place of the old ultra tactic *}
--- a/src/HOL/NSA/hypreal_arith.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/NSA/hypreal_arith.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -30,10 +30,10 @@
     Simplifier.simproc (the_context ())
       "fast_hypreal_arith" 
       ["(m::hypreal) < n", "(m::hypreal) <= n", "(m::hypreal) = n"]
-    (K LinArith.lin_arith_simproc);
+    (K Lin_Arith.lin_arith_simproc);
 
 val hypreal_arith_setup =
-  LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
+  Lin_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
    {add_mono_thms = add_mono_thms,
     mult_mono_thms = mult_mono_thms,
     inj_thms = real_inj_thms @ inj_thms,
--- a/src/HOL/Nat.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Nat.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -63,9 +63,8 @@
 end
 
 lemma Suc_not_Zero: "Suc m \<noteq> 0"
-  apply (simp add: Zero_nat_def Suc_def Abs_Nat_inject [unfolded mem_def]
+  by (simp add: Zero_nat_def Suc_def Abs_Nat_inject [unfolded mem_def]
     Rep_Nat [unfolded mem_def] Suc_RepI Zero_RepI Suc_Rep_not_Zero_Rep [unfolded mem_def])
-  done
 
 lemma Zero_not_Suc: "0 \<noteq> Suc m"
   by (rule not_sym, rule Suc_not_Zero not_sym)
@@ -82,7 +81,7 @@
   done
 
 lemma nat_induct [case_names 0 Suc, induct type: nat]:
-  -- {* for backward compatibility -- naming of variables differs *}
+  -- {* for backward compatibility -- names of variables differ *}
   fixes n
   assumes "P 0"
     and "\<And>n. P n \<Longrightarrow> P (Suc n)"
@@ -1345,19 +1344,13 @@
   shows "u = s"
   using 2 1 by (rule trans)
 
+setup Arith_Data.setup
+
 use "Tools/nat_arith.ML"
 declaration {* K Nat_Arith.setup *}
 
-ML{*
-structure ArithFacts =
-  NamedThmsFun(val name = "arith"
-               val description = "arith facts - only ground formulas");
-*}
-
-setup ArithFacts.setup
-
 use "Tools/lin_arith.ML"
-declaration {* K LinArith.setup *}
+declaration {* K Lin_Arith.setup *}
 
 lemmas [arith_split] = nat_diff_split split_min split_max
 
--- a/src/HOL/NatBin.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/NatBin.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -7,6 +7,7 @@
 
 theory NatBin
 imports IntDiv
+uses ("Tools/nat_simprocs.ML")
 begin
 
 text {*
@@ -40,9 +41,7 @@
 
 subsection {* Predicate for negative binary numbers *}
 
-definition
-  neg  :: "int \<Rightarrow> bool"
-where
+definition neg  :: "int \<Rightarrow> bool" where
   "neg Z \<longleftrightarrow> Z < 0"
 
 lemma not_neg_int [simp]: "~ neg (of_nat n)"
@@ -652,7 +651,7 @@
 val numeral_ss = @{simpset} addsimps @{thms numerals};
 
 val nat_bin_arith_setup =
- LinArith.map_data
+ Lin_Arith.map_data
    (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
      {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms,
       inj_thms = inj_thms,
@@ -818,4 +817,159 @@
      "(k*m) div (k*n) = (if k = (0::nat) then 0 else m div n)"
 by (simp add: nat_mult_div_cancel1)
 
+
+subsection {* Simprocs for the Naturals *}
+
+use "Tools/nat_simprocs.ML"
+declaration {* K nat_simprocs_setup *}
+
+subsubsection{*For simplifying @{term "Suc m - K"} and  @{term "K - Suc m"}*}
+
+text{*Where K above is a literal*}
+
+lemma Suc_diff_eq_diff_pred: "Numeral0 < n ==> Suc m - n = m - (n - Numeral1)"
+by (simp add: numeral_0_eq_0 numeral_1_eq_1 split add: nat_diff_split)
+
+text {*Now just instantiating @{text n} to @{text "number_of v"} does
+  the right simplification, but with some redundant inequality
+  tests.*}
+lemma neg_number_of_pred_iff_0:
+  "neg (number_of (Int.pred v)::int) = (number_of v = (0::nat))"
+apply (subgoal_tac "neg (number_of (Int.pred v)) = (number_of v < Suc 0) ")
+apply (simp only: less_Suc_eq_le le_0_eq)
+apply (subst less_number_of_Suc, simp)
+done
+
+text{*No longer required as a simprule because of the @{text inverse_fold}
+   simproc*}
+lemma Suc_diff_number_of:
+     "Int.Pls < v ==>
+      Suc m - (number_of v) = m - (number_of (Int.pred v))"
+apply (subst Suc_diff_eq_diff_pred)
+apply simp
+apply (simp del: nat_numeral_1_eq_1)
+apply (auto simp only: diff_nat_number_of less_0_number_of [symmetric]
+                        neg_number_of_pred_iff_0)
+done
+
+lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n"
+by (simp add: numerals split add: nat_diff_split)
+
+
+subsubsection{*For @{term nat_case} and @{term nat_rec}*}
+
+lemma nat_case_number_of [simp]:
+     "nat_case a f (number_of v) =
+        (let pv = number_of (Int.pred v) in
+         if neg pv then a else f (nat pv))"
+by (simp split add: nat.split add: Let_def neg_number_of_pred_iff_0)
+
+lemma nat_case_add_eq_if [simp]:
+     "nat_case a f ((number_of v) + n) =
+       (let pv = number_of (Int.pred v) in
+         if neg pv then nat_case a f n else f (nat pv + n))"
+apply (subst add_eq_if)
+apply (simp split add: nat.split
+            del: nat_numeral_1_eq_1
+            add: nat_numeral_1_eq_1 [symmetric]
+                 numeral_1_eq_Suc_0 [symmetric]
+                 neg_number_of_pred_iff_0)
+done
+
+lemma nat_rec_number_of [simp]:
+     "nat_rec a f (number_of v) =
+        (let pv = number_of (Int.pred v) in
+         if neg pv then a else f (nat pv) (nat_rec a f (nat pv)))"
+apply (case_tac " (number_of v) ::nat")
+apply (simp_all (no_asm_simp) add: Let_def neg_number_of_pred_iff_0)
+apply (simp split add: split_if_asm)
+done
+
+lemma nat_rec_add_eq_if [simp]:
+     "nat_rec a f (number_of v + n) =
+        (let pv = number_of (Int.pred v) in
+         if neg pv then nat_rec a f n
+                   else f (nat pv + n) (nat_rec a f (nat pv + n)))"
+apply (subst add_eq_if)
+apply (simp split add: nat.split
+            del: nat_numeral_1_eq_1
+            add: nat_numeral_1_eq_1 [symmetric]
+                 numeral_1_eq_Suc_0 [symmetric]
+                 neg_number_of_pred_iff_0)
+done
+
+
+subsubsection{*Various Other Lemmas*}
+
+text {*Evens and Odds, for Mutilated Chess Board*}
+
+text{*Lemmas for specialist use, NOT as default simprules*}
+lemma nat_mult_2: "2 * z = (z+z::nat)"
+proof -
+  have "2*z = (1 + 1)*z" by simp
+  also have "... = z+z" by (simp add: left_distrib)
+  finally show ?thesis .
+qed
+
+lemma nat_mult_2_right: "z * 2 = (z+z::nat)"
+by (subst mult_commute, rule nat_mult_2)
+
+text{*Case analysis on @{term "n<2"}*}
+lemma less_2_cases: "(n::nat) < 2 ==> n = 0 | n = Suc 0"
+by arith
+
+lemma div2_Suc_Suc [simp]: "Suc(Suc m) div 2 = Suc (m div 2)"
+by arith
+
+lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)"
+by (simp add: nat_mult_2 [symmetric])
+
+lemma mod2_Suc_Suc [simp]: "Suc(Suc(m)) mod 2 = m mod 2"
+apply (subgoal_tac "m mod 2 < 2")
+apply (erule less_2_cases [THEN disjE])
+apply (simp_all (no_asm_simp) add: Let_def mod_Suc nat_1)
+done
+
+lemma mod2_gr_0 [simp]: "!!m::nat. (0 < m mod 2) = (m mod 2 = 1)"
+apply (subgoal_tac "m mod 2 < 2")
+apply (force simp del: mod_less_divisor, simp)
+done
+
+text{*Removal of Small Numerals: 0, 1 and (in additive positions) 2*}
+
+lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)"
+by simp
+
+lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)"
+by simp
+
+text{*Can be used to eliminate long strings of Sucs, but not by default*}
+lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n"
+by simp
+
+
+text{*These lemmas collapse some needless occurrences of Suc:
+    at least three Sucs, since two and fewer are rewritten back to Suc again!
+    We already have some rules to simplify operands smaller than 3.*}
+
+lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)"
+by (simp add: Suc3_eq_add_3)
+
+lemma mod_Suc_eq_mod_add3 [simp]: "m mod (Suc (Suc (Suc n))) = m mod (3+n)"
+by (simp add: Suc3_eq_add_3)
+
+lemma Suc_div_eq_add3_div: "(Suc (Suc (Suc m))) div n = (3+m) div n"
+by (simp add: Suc3_eq_add_3)
+
+lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n"
+by (simp add: Suc3_eq_add_3)
+
+lemmas Suc_div_eq_add3_div_number_of =
+    Suc_div_eq_add3_div [of _ "number_of v", standard]
+declare Suc_div_eq_add3_div_number_of [simp]
+
+lemmas Suc_mod_eq_add3_mod_number_of =
+    Suc_mod_eq_add3_mod [of _ "number_of v", standard]
+declare Suc_mod_eq_add3_mod_number_of [simp]
+
 end
--- a/src/HOL/Nominal/nominal_atoms.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Nominal/nominal_atoms.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -106,8 +106,8 @@
               val ak_type = Type (Sign.intern_type thy1 ak,[])
               val ak_sign = Sign.intern_const thy1 ak 
               
-              val inj_type = @{typ nat}-->ak_type
-              val inj_on_type = inj_type-->(@{typ "nat set"})-->@{typ bool}  
+              val inj_type = @{typ nat} --> ak_type
+              val inj_on_type = inj_type --> @{typ "nat set"} --> @{typ bool}
 
               (* first statement *)
               val stmnt1 = HOLogic.mk_Trueprop 
--- a/src/HOL/OrderedGroup.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/OrderedGroup.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -316,6 +316,9 @@
 lemma eq_iff_diff_eq_0: "a = b \<longleftrightarrow> a - b = 0"
 by (simp add: algebra_simps)
 
+lemma diff_eq_0_iff_eq [simp, noatp]: "a - b = 0 \<longleftrightarrow> a = b"
+by (simp add: algebra_simps)
+
 end
 
 subsection {* (Partially) Ordered Groups *} 
@@ -463,7 +466,7 @@
   then show ?thesis by simp
 qed
 
-lemma add_neg_nonpos: 
+lemma add_neg_nonpos:
   assumes "a < 0" and "b \<le> 0" shows "a + b < 0"
 proof -
   have "a + b < 0 + 0"
@@ -491,6 +494,10 @@
   then show ?thesis by simp
 qed
 
+lemmas add_sign_intros =
+  add_pos_nonneg add_pos_pos add_nonneg_pos add_nonneg_nonneg
+  add_neg_nonpos add_neg_neg add_nonpos_neg add_nonpos_nonpos
+
 lemma add_nonneg_eq_0_iff:
   assumes x: "0 \<le> x" and y: "0 \<le> y"
   shows "x + y = 0 \<longleftrightarrow> x = 0 \<and> y = 0"
@@ -1296,7 +1303,7 @@
 done
 
 lemma eq_eqI: "(x::'a::ab_group_add) - y = x' - y' \<Longrightarrow> (x = y) = (x' = y')"
-by (simp add: eq_iff_diff_eq_0[of x y] eq_iff_diff_eq_0[of x' y'])
+by (simp only: eq_iff_diff_eq_0[of x y] eq_iff_diff_eq_0[of x' y'])
 
 lemma diff_def: "(x::'a::ab_group_add) - y == x + (-y)"
 by (simp add: diff_minus)
@@ -1344,7 +1351,6 @@
 
 text{*Simplification of @{term "x-y < 0"}, etc.*}
 lemmas diff_less_0_iff_less [simp, noatp] = less_iff_diff_less_0 [symmetric]
-lemmas diff_eq_0_iff_eq [simp, noatp] = eq_iff_diff_eq_0 [symmetric]
 lemmas diff_le_0_iff_le [simp, noatp] = le_iff_diff_le_0 [symmetric]
 
 ML {*
--- a/src/HOL/Orderings.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Orderings.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -372,13 +372,14 @@
   Thm.declaration_attribute
     (fn _ => Data.map (AList.delete struct_eq s));
 
-val attribute =
-  Scan.lift ((Args.add -- Args.name >> (fn (_, s) => SOME s) ||
-      Args.del >> K NONE) --| Args.colon (* FIXME ||
-    Scan.succeed true *) ) -- Scan.lift Args.name --
-  Scan.repeat Args.term
-  >> (fn ((SOME tag, n), ts) => add_struct_thm (n, ts) tag
-       | ((NONE, n), ts) => del_struct (n, ts));
+val attrib_setup =
+  Attrib.setup @{binding order}
+    (Scan.lift ((Args.add -- Args.name >> (fn (_, s) => SOME s) || Args.del >> K NONE) --|
+      Args.colon (* FIXME || Scan.succeed true *) ) -- Scan.lift Args.name --
+      Scan.repeat Args.term
+      >> (fn ((SOME tag, n), ts) => add_struct_thm (n, ts) tag
+           | ((NONE, n), ts) => del_struct (n, ts)))
+    "theorems controlling transitivity reasoner";
 
 
 (** Diagnostic command **)
@@ -397,8 +398,9 @@
 (** Setup **)
 
 val setup =
-  Method.setup @{binding order} (Scan.succeed (SIMPLE_METHOD' o order_tac [])) "transitivity reasoner" #>
-  Attrib.setup @{binding order} attribute "theorems controlling transitivity reasoner";
+  Method.setup @{binding order} (Scan.succeed (SIMPLE_METHOD' o order_tac []))
+    "transitivity reasoner" #>
+  attrib_setup;
 
 end;
 
--- a/src/HOL/Parity.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Parity.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -5,7 +5,7 @@
 header {* Even and Odd for int and nat *}
 
 theory Parity
-imports Plain Presburger Main
+imports Main
 begin
 
 class even_odd = 
--- a/src/HOL/Power.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Power.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -186,6 +186,10 @@
 apply (auto simp add: abs_mult)
 done
 
+lemma abs_power_minus [simp]:
+  fixes a:: "'a::{ordered_idom,recpower}" shows "abs((-a) ^ n) = abs(a ^ n)"
+  by (simp add: abs_minus_cancel power_abs) 
+
 lemma zero_less_power_abs_iff [simp,noatp]:
      "(0 < (abs a)^n) = (a \<noteq> (0::'a::{ordered_idom,recpower}) | n=0)"
 proof (induct "n")
--- a/src/HOL/Presburger.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Presburger.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -7,6 +7,7 @@
 theory Presburger
 imports Groebner_Basis SetInterval
 uses
+  "Tools/Qelim/qelim.ML"
   "Tools/Qelim/cooper_data.ML"
   "Tools/Qelim/generated_cooper.ML"
   ("Tools/Qelim/cooper.ML")
@@ -438,12 +439,7 @@
 
 use "Tools/Qelim/presburger.ML"
 
-declaration {* fn _ =>
-  arith_tactic_add
-    (mk_arith_tactic "presburger" (fn ctxt => fn i => fn st =>
-       (warning "Trying Presburger arithmetic ...";   
-    Presburger.cooper_tac true [] [] ctxt i st)))
-*}
+setup {* Arith_Data.add_tactic "Presburger arithmetic" (K (Presburger.cooper_tac true [] [])) *}
 
 method_setup presburger = {*
 let
--- a/src/HOL/Product_Type.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Product_Type.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -981,7 +981,8 @@
   | strip_abs_split i (u as Const ("split", _) $ t) = (case strip_abs_split (i+1) t of
         (v :: v' :: vs, u) => (HOLogic.mk_prod (v, v') :: vs, u)
       | _ => ([], u))
-  | strip_abs_split i t = ([], t);
+  | strip_abs_split i t =
+      strip_abs_split i (Abs ("x", hd (binder_types (fastype_of t)), t $ Bound 0));
 
 fun let_codegen thy defs dep thyname brack t gr = (case strip_comb t of
     (t1 as Const ("Let", _), t2 :: t3 :: ts) =>
@@ -1018,19 +1019,17 @@
 
 fun split_codegen thy defs dep thyname brack t gr = (case strip_comb t of
     (t1 as Const ("split", _), t2 :: ts) =>
-      (case strip_abs_split 1 (t1 $ t2) of
-         ([p], u) =>
-           let
-             val (q, gr1) = Codegen.invoke_codegen thy defs dep thyname false p gr;
-             val (pu, gr2) = Codegen.invoke_codegen thy defs dep thyname false u gr1;
-             val (pargs, gr3) = fold_map
-               (Codegen.invoke_codegen thy defs dep thyname true) ts gr2
-           in
-             SOME (Codegen.mk_app brack
-               (Pretty.block [Codegen.str "(fn ", q, Codegen.str " =>",
-                 Pretty.brk 1, pu, Codegen.str ")"]) pargs, gr2)
-           end
-       | _ => NONE)
+      let
+        val ([p], u) = strip_abs_split 1 (t1 $ t2);
+        val (q, gr1) = Codegen.invoke_codegen thy defs dep thyname false p gr;
+        val (pu, gr2) = Codegen.invoke_codegen thy defs dep thyname false u gr1;
+        val (pargs, gr3) = fold_map
+          (Codegen.invoke_codegen thy defs dep thyname true) ts gr2
+      in
+        SOME (Codegen.mk_app brack
+          (Pretty.block [Codegen.str "(fn ", q, Codegen.str " =>",
+            Pretty.brk 1, pu, Codegen.str ")"]) pargs, gr2)
+      end
   | _ => NONE);
 
 in
--- a/src/HOL/Rational.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Rational.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -691,7 +691,6 @@
     \<longleftrightarrow> Fract a b < Fract c d"
     using not_zero `b * d > 0`
     by (simp add: of_rat_rat of_int_divide_less_eq of_int_mult [symmetric] del: of_int_mult)
-      (auto intro: mult_strict_right_mono mult_right_less_imp_less)
 qed
 
 lemma of_rat_less_eq:
--- a/src/HOL/RealVector.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/RealVector.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -145,7 +145,7 @@
   and scaleR_scaleR: "scaleR a (scaleR b x) = scaleR (a * b) x"
   and scaleR_one: "scaleR 1 x = x"
 
-interpretation real_vector!:
+interpretation real_vector:
   vector_space "scaleR :: real \<Rightarrow> 'a \<Rightarrow> 'a::real_vector"
 apply unfold_locales
 apply (rule scaleR_right_distrib)
@@ -190,10 +190,10 @@
 
 end
 
-interpretation scaleR_left!: additive "(\<lambda>a. scaleR a x::'a::real_vector)"
+interpretation scaleR_left: additive "(\<lambda>a. scaleR a x::'a::real_vector)"
 proof qed (rule scaleR_left_distrib)
 
-interpretation scaleR_right!: additive "(\<lambda>x. scaleR a x::'a::real_vector)"
+interpretation scaleR_right: additive "(\<lambda>x. scaleR a x::'a::real_vector)"
 proof qed (rule scaleR_right_distrib)
 
 lemma nonzero_inverse_scaleR_distrib:
@@ -789,7 +789,7 @@
 
 end
 
-interpretation mult!:
+interpretation mult:
   bounded_bilinear "op * :: 'a \<Rightarrow> 'a \<Rightarrow> 'a::real_normed_algebra"
 apply (rule bounded_bilinear.intro)
 apply (rule left_distrib)
@@ -800,19 +800,19 @@
 apply (simp add: norm_mult_ineq)
 done
 
-interpretation mult_left!:
+interpretation mult_left:
   bounded_linear "(\<lambda>x::'a::real_normed_algebra. x * y)"
 by (rule mult.bounded_linear_left)
 
-interpretation mult_right!:
+interpretation mult_right:
   bounded_linear "(\<lambda>y::'a::real_normed_algebra. x * y)"
 by (rule mult.bounded_linear_right)
 
-interpretation divide!:
+interpretation divide:
   bounded_linear "(\<lambda>x::'a::real_normed_field. x / y)"
 unfolding divide_inverse by (rule mult.bounded_linear_left)
 
-interpretation scaleR!: bounded_bilinear "scaleR"
+interpretation scaleR: bounded_bilinear "scaleR"
 apply (rule bounded_bilinear.intro)
 apply (rule scaleR_left_distrib)
 apply (rule scaleR_right_distrib)
@@ -822,13 +822,13 @@
 apply (simp add: norm_scaleR)
 done
 
-interpretation scaleR_left!: bounded_linear "\<lambda>r. scaleR r x"
+interpretation scaleR_left: bounded_linear "\<lambda>r. scaleR r x"
 by (rule scaleR.bounded_linear_left)
 
-interpretation scaleR_right!: bounded_linear "\<lambda>x. scaleR r x"
+interpretation scaleR_right: bounded_linear "\<lambda>x. scaleR r x"
 by (rule scaleR.bounded_linear_right)
 
-interpretation of_real!: bounded_linear "\<lambda>r. of_real r"
+interpretation of_real: bounded_linear "\<lambda>r. of_real r"
 unfolding of_real_def by (rule scaleR.bounded_linear_left)
 
 end
--- a/src/HOL/Ring_and_Field.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Ring_and_Field.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -534,7 +534,156 @@
 by (simp add: divide_inverse)
 
 lemma add_divide_distrib: "(a+b) / c = a/c + b/c"
-by (simp add: divide_inverse algebra_simps) 
+by (simp add: divide_inverse algebra_simps)
+
+text{*There is no slick version using division by zero.*}
+lemma inverse_add:
+  "[| a \<noteq> 0;  b \<noteq> 0 |]
+   ==> inverse a + inverse b = (a + b) * inverse a * inverse b"
+by (simp add: division_ring_inverse_add mult_ac)
+
+lemma nonzero_mult_divide_mult_cancel_left [simp, noatp]:
+assumes [simp]: "b\<noteq>0" and [simp]: "c\<noteq>0" shows "(c*a)/(c*b) = a/b"
+proof -
+  have "(c*a)/(c*b) = c * a * (inverse b * inverse c)"
+    by (simp add: divide_inverse nonzero_inverse_mult_distrib)
+  also have "... =  a * inverse b * (inverse c * c)"
+    by (simp only: mult_ac)
+  also have "... =  a * inverse b" by simp
+    finally show ?thesis by (simp add: divide_inverse)
+qed
+
+lemma nonzero_mult_divide_mult_cancel_right [simp, noatp]:
+  "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (a * c) / (b * c) = a / b"
+by (simp add: mult_commute [of _ c])
+
+lemma divide_1 [simp]: "a / 1 = a"
+by (simp add: divide_inverse)
+
+lemma times_divide_eq_right: "a * (b / c) = (a * b) / c"
+by (simp add: divide_inverse mult_assoc)
+
+lemma times_divide_eq_left: "(b / c) * a = (b * a) / c"
+by (simp add: divide_inverse mult_ac)
+
+text {* These are later declared as simp rules. *}
+lemmas times_divide_eq [noatp] = times_divide_eq_right times_divide_eq_left
+
+lemma add_frac_eq:
+  assumes "y \<noteq> 0" and "z \<noteq> 0"
+  shows "x / y + w / z = (x * z + w * y) / (y * z)"
+proof -
+  have "x / y + w / z = (x * z) / (y * z) + (y * w) / (y * z)"
+    using assms by simp
+  also have "\<dots> = (x * z + y * w) / (y * z)"
+    by (simp only: add_divide_distrib)
+  finally show ?thesis
+    by (simp only: mult_commute)
+qed
+
+text{*Special Cancellation Simprules for Division*}
+
+lemma nonzero_mult_divide_cancel_right [simp, noatp]:
+  "b \<noteq> 0 \<Longrightarrow> a * b / b = a"
+using nonzero_mult_divide_mult_cancel_right [of 1 b a] by simp
+
+lemma nonzero_mult_divide_cancel_left [simp, noatp]:
+  "a \<noteq> 0 \<Longrightarrow> a * b / a = b"
+using nonzero_mult_divide_mult_cancel_left [of 1 a b] by simp
+
+lemma nonzero_divide_mult_cancel_right [simp, noatp]:
+  "\<lbrakk>a \<noteq> 0; b \<noteq> 0\<rbrakk> \<Longrightarrow> b / (a * b) = 1 / a"
+using nonzero_mult_divide_mult_cancel_right [of a b 1] by simp
+
+lemma nonzero_divide_mult_cancel_left [simp, noatp]:
+  "\<lbrakk>a \<noteq> 0; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / (a * b) = 1 / b"
+using nonzero_mult_divide_mult_cancel_left [of b a 1] by simp
+
+lemma nonzero_mult_divide_mult_cancel_left2 [simp, noatp]:
+  "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (c * a) / (b * c) = a / b"
+using nonzero_mult_divide_mult_cancel_left [of b c a] by (simp add: mult_ac)
+
+lemma nonzero_mult_divide_mult_cancel_right2 [simp, noatp]:
+  "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (a * c) / (c * b) = a / b"
+using nonzero_mult_divide_mult_cancel_right [of b c a] by (simp add: mult_ac)
+
+lemma minus_divide_left: "- (a / b) = (-a) / b"
+by (simp add: divide_inverse)
+
+lemma nonzero_minus_divide_right: "b \<noteq> 0 ==> - (a / b) = a / (- b)"
+by (simp add: divide_inverse nonzero_inverse_minus_eq)
+
+lemma nonzero_minus_divide_divide: "b \<noteq> 0 ==> (-a) / (-b) = a / b"
+by (simp add: divide_inverse nonzero_inverse_minus_eq)
+
+lemma divide_minus_left [simp, noatp]: "(-a) / b = - (a / b)"
+by (simp add: divide_inverse)
+
+lemma diff_divide_distrib: "(a - b) / c = a / c - b / c"
+by (simp add: diff_minus add_divide_distrib)
+
+lemma add_divide_eq_iff:
+  "z \<noteq> 0 \<Longrightarrow> x + y / z = (z * x + y) / z"
+by (simp add: add_divide_distrib)
+
+lemma divide_add_eq_iff:
+  "z \<noteq> 0 \<Longrightarrow> x / z + y = (x + z * y) / z"
+by (simp add: add_divide_distrib)
+
+lemma diff_divide_eq_iff:
+  "z \<noteq> 0 \<Longrightarrow> x - y / z = (z * x - y) / z"
+by (simp add: diff_divide_distrib)
+
+lemma divide_diff_eq_iff:
+  "z \<noteq> 0 \<Longrightarrow> x / z - y = (x - z * y) / z"
+by (simp add: diff_divide_distrib)
+
+lemma nonzero_eq_divide_eq: "c \<noteq> 0 \<Longrightarrow> a = b / c \<longleftrightarrow> a * c = b"
+proof -
+  assume [simp]: "c \<noteq> 0"
+  have "a = b / c \<longleftrightarrow> a * c = (b / c) * c" by simp
+  also have "... \<longleftrightarrow> a * c = b" by (simp add: divide_inverse mult_assoc)
+  finally show ?thesis .
+qed
+
+lemma nonzero_divide_eq_eq: "c \<noteq> 0 \<Longrightarrow> b / c = a \<longleftrightarrow> b = a * c"
+proof -
+  assume [simp]: "c \<noteq> 0"
+  have "b / c = a \<longleftrightarrow> (b / c) * c = a * c" by simp
+  also have "... \<longleftrightarrow> b = a * c" by (simp add: divide_inverse mult_assoc) 
+  finally show ?thesis .
+qed
+
+lemma divide_eq_imp: "c \<noteq> 0 \<Longrightarrow> b = a * c \<Longrightarrow> b / c = a"
+by simp
+
+lemma eq_divide_imp: "c \<noteq> 0 \<Longrightarrow> a * c = b \<Longrightarrow> a = b / c"
+by (erule subst, simp)
+
+lemmas field_eq_simps[noatp] = algebra_simps
+  (* pull / out*)
+  add_divide_eq_iff divide_add_eq_iff
+  diff_divide_eq_iff divide_diff_eq_iff
+  (* multiply eqn *)
+  nonzero_eq_divide_eq nonzero_divide_eq_eq
+(* is added later:
+  times_divide_eq_left times_divide_eq_right
+*)
+
+text{*An example:*}
+lemma "\<lbrakk>a\<noteq>b; c\<noteq>d; e\<noteq>f\<rbrakk> \<Longrightarrow> ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) = 1"
+apply(subgoal_tac "(c-d)*(e-f)*(a-b) \<noteq> 0")
+ apply(simp add:field_eq_simps)
+apply(simp)
+done
+
+lemma diff_frac_eq:
+  "y \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> x / y - w / z = (x * z - w * y) / (y * z)"
+by (simp add: field_eq_simps times_divide_eq)
+
+lemma frac_eq_eq:
+  "y \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> (x / y = w / z) = (x * z = w * y)"
+by (simp add: field_eq_simps times_divide_eq)
 
 end
 
@@ -580,11 +729,15 @@
 subclass pordered_semiring ..
 
 lemma mult_nonneg_nonneg: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a * b"
-by (drule mult_left_mono [of zero b], auto)
+using mult_left_mono [of zero b a] by simp
 
 lemma mult_nonneg_nonpos: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> a * b \<le> 0"
-by (drule mult_left_mono [of b zero], auto)
-
+using mult_left_mono [of b zero a] by simp
+
+lemma mult_nonpos_nonneg: "a \<le> 0 \<Longrightarrow> 0 \<le> b \<Longrightarrow> a * b \<le> 0"
+using mult_right_mono [of a zero b] by simp
+
+text {* Legacy - use @{text mult_nonpos_nonneg} *}
 lemma mult_nonneg_nonpos2: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> b * a \<le> 0" 
 by (drule mult_right_mono [of b zero], auto)
 
@@ -637,31 +790,32 @@
   "a * c \<le> b * c \<Longrightarrow> 0 < c \<Longrightarrow> a \<le> b"
 by (force simp add: mult_strict_right_mono not_less [symmetric])
 
-lemma mult_pos_pos:
-  "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> 0 < a * b"
-by (drule mult_strict_left_mono [of zero b], auto)
-
-lemma mult_pos_neg:
-  "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> a * b < 0"
-by (drule mult_strict_left_mono [of b zero], auto)
-
-lemma mult_pos_neg2:
-  "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> b * a < 0" 
+lemma mult_pos_pos: "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> 0 < a * b"
+using mult_strict_left_mono [of zero b a] by simp
+
+lemma mult_pos_neg: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> a * b < 0"
+using mult_strict_left_mono [of b zero a] by simp
+
+lemma mult_neg_pos: "a < 0 \<Longrightarrow> 0 < b \<Longrightarrow> a * b < 0"
+using mult_strict_right_mono [of a zero b] by simp
+
+text {* Legacy - use @{text mult_neg_pos} *}
+lemma mult_pos_neg2: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> b * a < 0" 
 by (drule mult_strict_right_mono [of b zero], auto)
 
 lemma zero_less_mult_pos:
   "0 < a * b \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
-apply (cases "b\<le>0") 
+apply (cases "b\<le>0")
  apply (auto simp add: le_less not_less)
-apply (drule_tac mult_pos_neg [of a b]) 
+apply (drule_tac mult_pos_neg [of a b])
  apply (auto dest: less_not_sym)
 done
 
 lemma zero_less_mult_pos2:
   "0 < b * a \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
-apply (cases "b\<le>0") 
+apply (cases "b\<le>0")
  apply (auto simp add: le_less not_less)
-apply (drule_tac mult_pos_neg2 [of a b]) 
+apply (drule_tac mult_pos_neg2 [of a b])
  apply (auto dest: less_not_sym)
 done
 
@@ -670,9 +824,9 @@
   assumes "a < b" and "c < d" and "0 < b" and "0 \<le> c"
   shows "a * c < b * d"
   using assms apply (cases "c=0")
-  apply (simp add: mult_pos_pos) 
+  apply (simp add: mult_pos_pos)
   apply (erule mult_strict_right_mono [THEN less_trans])
-  apply (force simp add: le_less) 
+  apply (force simp add: le_less)
   apply (erule mult_strict_left_mono, assumption)
   done
 
@@ -811,9 +965,8 @@
   apply (simp_all add: minus_mult_right [symmetric]) 
   done
 
-lemma mult_nonpos_nonpos:
-  "a \<le> 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> 0 \<le> a * b"
-by (drule mult_right_mono_neg [of a zero b]) auto
+lemma mult_nonpos_nonpos: "a \<le> 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> 0 \<le> a * b"
+using mult_right_mono_neg [of a zero b] by simp
 
 lemma split_mult_pos_le:
   "(0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0) \<Longrightarrow> 0 \<le> a * b"
@@ -857,21 +1010,14 @@
 
 subclass ordered_ring ..
 
-lemma mult_strict_left_mono_neg:
-  "b < a \<Longrightarrow> c < 0 \<Longrightarrow> c * a < c * b"
-  apply (drule mult_strict_left_mono [of _ _ "uminus c"])
-  apply (simp_all add: minus_mult_left [symmetric]) 
-  done
-
-lemma mult_strict_right_mono_neg:
-  "b < a \<Longrightarrow> c < 0 \<Longrightarrow> a * c < b * c"
-  apply (drule mult_strict_right_mono [of _ _ "uminus c"])
-  apply (simp_all add: minus_mult_right [symmetric]) 
-  done
-
-lemma mult_neg_neg:
-  "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> 0 < a * b"
-by (drule mult_strict_right_mono_neg, auto)
+lemma mult_strict_left_mono_neg: "b < a \<Longrightarrow> c < 0 \<Longrightarrow> c * a < c * b"
+using mult_strict_left_mono [of b a "- c"] by simp
+
+lemma mult_strict_right_mono_neg: "b < a \<Longrightarrow> c < 0 \<Longrightarrow> a * c < b * c"
+using mult_strict_right_mono [of b a "- c"] by simp
+
+lemma mult_neg_neg: "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> 0 < a * b"
+using mult_strict_right_mono_neg [of a zero b] by simp
 
 subclass ring_no_zero_divisors
 proof
@@ -987,11 +1133,32 @@
   "c * a \<le> c * b \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
 by (simp add: not_less [symmetric] mult_less_cancel_left_disj)
 
+lemma mult_le_cancel_left_pos:
+  "0 < c \<Longrightarrow> c * a \<le> c * b \<longleftrightarrow> a \<le> b"
+by (auto simp: mult_le_cancel_left)
+
+lemma mult_le_cancel_left_neg:
+  "c < 0 \<Longrightarrow> c * a \<le> c * b \<longleftrightarrow> b \<le> a"
+by (auto simp: mult_le_cancel_left)
+
+lemma mult_less_cancel_left_pos:
+  "0 < c \<Longrightarrow> c * a < c * b \<longleftrightarrow> a < b"
+by (auto simp: mult_less_cancel_left)
+
+lemma mult_less_cancel_left_neg:
+  "c < 0 \<Longrightarrow> c * a < c * b \<longleftrightarrow> b < a"
+by (auto simp: mult_less_cancel_left)
+
 end
 
 text{*Legacy - use @{text algebra_simps} *}
 lemmas ring_simps[noatp] = algebra_simps
 
+lemmas mult_sign_intros =
+  mult_nonneg_nonneg mult_nonneg_nonpos
+  mult_nonpos_nonneg mult_nonpos_nonpos
+  mult_pos_pos mult_pos_neg
+  mult_neg_pos mult_neg_neg
 
 class pordered_comm_ring = comm_ring + pordered_comm_semiring
 begin
@@ -1191,12 +1358,6 @@
     thus ?thesis by force
   qed
 
-text{*There is no slick version using division by zero.*}
-lemma inverse_add:
-  "[|a \<noteq> 0;  b \<noteq> 0|]
-   ==> inverse a + inverse b = (a+b) * inverse a * inverse (b::'a::field)"
-by (simp add: division_ring_inverse_add mult_ac)
-
 lemma inverse_divide [simp]:
   "inverse (a/b) = b / (a::'a::{field,division_by_zero})"
 by (simp add: divide_inverse mult_commute)
@@ -1208,44 +1369,18 @@
 field} but none for class @{text field} and @{text nonzero_divides}
 because the latter are covered by a simproc. *}
 
-lemma nonzero_mult_divide_mult_cancel_left[simp,noatp]:
-assumes [simp]: "b\<noteq>0" and [simp]: "c\<noteq>0" shows "(c*a)/(c*b) = a/(b::'a::field)"
-proof -
-  have "(c*a)/(c*b) = c * a * (inverse b * inverse c)"
-    by (simp add: divide_inverse nonzero_inverse_mult_distrib)
-  also have "... =  a * inverse b * (inverse c * c)"
-    by (simp only: mult_ac)
-  also have "... =  a * inverse b" by simp
-    finally show ?thesis by (simp add: divide_inverse)
-qed
-
 lemma mult_divide_mult_cancel_left:
   "c\<noteq>0 ==> (c*a) / (c*b) = a / (b::'a::{field,division_by_zero})"
 apply (cases "b = 0")
 apply (simp_all add: nonzero_mult_divide_mult_cancel_left)
 done
 
-lemma nonzero_mult_divide_mult_cancel_right [noatp]:
-  "[|b\<noteq>0; c\<noteq>0|] ==> (a*c) / (b*c) = a/(b::'a::field)"
-by (simp add: mult_commute [of _ c] nonzero_mult_divide_mult_cancel_left) 
-
 lemma mult_divide_mult_cancel_right:
   "c\<noteq>0 ==> (a*c) / (b*c) = a / (b::'a::{field,division_by_zero})"
 apply (cases "b = 0")
 apply (simp_all add: nonzero_mult_divide_mult_cancel_right)
 done
 
-lemma divide_1 [simp]: "a/1 = (a::'a::field)"
-by (simp add: divide_inverse)
-
-lemma times_divide_eq_right: "a * (b/c) = (a*b) / (c::'a::field)"
-by (simp add: divide_inverse mult_assoc)
-
-lemma times_divide_eq_left: "(b/c) * a = (b*a) / (c::'a::field)"
-by (simp add: divide_inverse mult_ac)
-
-lemmas times_divide_eq[noatp] = times_divide_eq_right times_divide_eq_left
-
 lemma divide_divide_eq_right [simp,noatp]:
   "a / (b/c) = (a*c) / (b::'a::{field,division_by_zero})"
 by (simp add: divide_inverse mult_ac)
@@ -1254,20 +1389,6 @@
   "(a / b) / (c::'a::{field,division_by_zero}) = a / (b*c)"
 by (simp add: divide_inverse mult_assoc)
 
-lemma add_frac_eq: "(y::'a::field) ~= 0 ==> z ~= 0 ==>
-    x / y + w / z = (x * z + w * y) / (y * z)"
-apply (subgoal_tac "x / y = (x * z) / (y * z)")
-apply (erule ssubst)
-apply (subgoal_tac "w / z = (w * y) / (y * z)")
-apply (erule ssubst)
-apply (rule add_divide_distrib [THEN sym])
-apply (subst mult_commute)
-apply (erule nonzero_mult_divide_mult_cancel_left [THEN sym])
-apply assumption
-apply (erule nonzero_mult_divide_mult_cancel_right [THEN sym])
-apply assumption
-done
-
 
 subsubsection{*Special Cancellation Simprules for Division*}
 
@@ -1276,140 +1397,29 @@
 shows "(c*a) / (c*b) = (if c=0 then 0 else a/b)"
 by (simp add: mult_divide_mult_cancel_left)
 
-lemma nonzero_mult_divide_cancel_right[simp,noatp]:
-  "b \<noteq> 0 \<Longrightarrow> a * b / b = (a::'a::field)"
-using nonzero_mult_divide_mult_cancel_right[of 1 b a] by simp
-
-lemma nonzero_mult_divide_cancel_left[simp,noatp]:
-  "a \<noteq> 0 \<Longrightarrow> a * b / a = (b::'a::field)"
-using nonzero_mult_divide_mult_cancel_left[of 1 a b] by simp
-
-
-lemma nonzero_divide_mult_cancel_right[simp,noatp]:
-  "\<lbrakk> a\<noteq>0; b\<noteq>0 \<rbrakk> \<Longrightarrow> b / (a * b) = 1/(a::'a::field)"
-using nonzero_mult_divide_mult_cancel_right[of a b 1] by simp
-
-lemma nonzero_divide_mult_cancel_left[simp,noatp]:
-  "\<lbrakk> a\<noteq>0; b\<noteq>0 \<rbrakk> \<Longrightarrow> a / (a * b) = 1/(b::'a::field)"
-using nonzero_mult_divide_mult_cancel_left[of b a 1] by simp
-
-
-lemma nonzero_mult_divide_mult_cancel_left2[simp,noatp]:
-  "[|b\<noteq>0; c\<noteq>0|] ==> (c*a) / (b*c) = a/(b::'a::field)"
-using nonzero_mult_divide_mult_cancel_left[of b c a] by(simp add:mult_ac)
-
-lemma nonzero_mult_divide_mult_cancel_right2[simp,noatp]:
-  "[|b\<noteq>0; c\<noteq>0|] ==> (a*c) / (c*b) = a/(b::'a::field)"
-using nonzero_mult_divide_mult_cancel_right[of b c a] by(simp add:mult_ac)
-
 
 subsection {* Division and Unary Minus *}
 
-lemma nonzero_minus_divide_left: "b \<noteq> 0 ==> - (a/b) = (-a) / (b::'a::field)"
-by (simp add: divide_inverse)
-
-lemma nonzero_minus_divide_right: "b \<noteq> 0 ==> - (a/b) = a / -(b::'a::field)"
-by (simp add: divide_inverse nonzero_inverse_minus_eq)
-
-lemma nonzero_minus_divide_divide: "b \<noteq> 0 ==> (-a)/(-b) = a / (b::'a::field)"
-by (simp add: divide_inverse nonzero_inverse_minus_eq)
-
-lemma minus_divide_left: "- (a/b) = (-a) / (b::'a::field)"
-by (simp add: divide_inverse)
-
 lemma minus_divide_right: "- (a/b) = a / -(b::'a::{field,division_by_zero})"
 by (simp add: divide_inverse)
 
-
-text{*The effect is to extract signs from divisions*}
-lemmas divide_minus_left[noatp] = minus_divide_left [symmetric]
-lemmas divide_minus_right[noatp] = minus_divide_right [symmetric]
-declare divide_minus_left [simp]   divide_minus_right [simp]
-
-lemma minus_divide_divide [simp]:
+lemma divide_minus_right [simp, noatp]:
+  "a / -(b::'a::{field,division_by_zero}) = -(a / b)"
+by (simp add: divide_inverse)
+
+lemma minus_divide_divide:
   "(-a)/(-b) = a / (b::'a::{field,division_by_zero})"
 apply (cases "b=0", simp) 
 apply (simp add: nonzero_minus_divide_divide) 
 done
 
-lemma diff_divide_distrib: "(a-b)/(c::'a::field) = a/c - b/c"
-by (simp add: diff_minus add_divide_distrib) 
-
-lemma add_divide_eq_iff:
-  "(z::'a::field) \<noteq> 0 \<Longrightarrow> x + y/z = (z*x + y)/z"
-by(simp add:add_divide_distrib nonzero_mult_divide_cancel_left)
-
-lemma divide_add_eq_iff:
-  "(z::'a::field) \<noteq> 0 \<Longrightarrow> x/z + y = (x + z*y)/z"
-by(simp add:add_divide_distrib nonzero_mult_divide_cancel_left)
-
-lemma diff_divide_eq_iff:
-  "(z::'a::field) \<noteq> 0 \<Longrightarrow> x - y/z = (z*x - y)/z"
-by(simp add:diff_divide_distrib nonzero_mult_divide_cancel_left)
-
-lemma divide_diff_eq_iff:
-  "(z::'a::field) \<noteq> 0 \<Longrightarrow> x/z - y = (x - z*y)/z"
-by(simp add:diff_divide_distrib nonzero_mult_divide_cancel_left)
-
-lemma nonzero_eq_divide_eq: "c\<noteq>0 ==> ((a::'a::field) = b/c) = (a*c = b)"
-proof -
-  assume [simp]: "c\<noteq>0"
-  have "(a = b/c) = (a*c = (b/c)*c)" by simp
-  also have "... = (a*c = b)" by (simp add: divide_inverse mult_assoc)
-  finally show ?thesis .
-qed
-
-lemma nonzero_divide_eq_eq: "c\<noteq>0 ==> (b/c = (a::'a::field)) = (b = a*c)"
-proof -
-  assume [simp]: "c\<noteq>0"
-  have "(b/c = a) = ((b/c)*c = a*c)"  by simp
-  also have "... = (b = a*c)"  by (simp add: divide_inverse mult_assoc) 
-  finally show ?thesis .
-qed
-
 lemma eq_divide_eq:
   "((a::'a::{field,division_by_zero}) = b/c) = (if c\<noteq>0 then a*c = b else a=0)"
-by (simp add: nonzero_eq_divide_eq) 
+by (simp add: nonzero_eq_divide_eq)
 
 lemma divide_eq_eq:
   "(b/c = (a::'a::{field,division_by_zero})) = (if c\<noteq>0 then b = a*c else a=0)"
-by (force simp add: nonzero_divide_eq_eq) 
-
-lemma divide_eq_imp: "(c::'a::{division_by_zero,field}) ~= 0 ==>
-    b = a * c ==> b / c = a"
-by (subst divide_eq_eq, simp)
-
-lemma eq_divide_imp: "(c::'a::{division_by_zero,field}) ~= 0 ==>
-    a * c = b ==> a = b / c"
-by (subst eq_divide_eq, simp)
-
-
-lemmas field_eq_simps[noatp] = algebra_simps
-  (* pull / out*)
-  add_divide_eq_iff divide_add_eq_iff
-  diff_divide_eq_iff divide_diff_eq_iff
-  (* multiply eqn *)
-  nonzero_eq_divide_eq nonzero_divide_eq_eq
-(* is added later:
-  times_divide_eq_left times_divide_eq_right
-*)
-
-text{*An example:*}
-lemma fixes a b c d e f :: "'a::field"
-shows "\<lbrakk>a\<noteq>b; c\<noteq>d; e\<noteq>f \<rbrakk> \<Longrightarrow> ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) = 1"
-apply(subgoal_tac "(c-d)*(e-f)*(a-b) \<noteq> 0")
- apply(simp add:field_eq_simps)
-apply(simp)
-done
-
-
-lemma diff_frac_eq: "(y::'a::field) ~= 0 ==> z ~= 0 ==>
-    x / y - w / z = (x * z - w * y) / (y * z)"
-by (simp add:field_eq_simps times_divide_eq)
-
-lemma frac_eq_eq: "(y::'a::field) ~= 0 ==> z ~= 0 ==>
-    (x / y = w / z) = (x * z = w * y)"
-by (simp add:field_eq_simps times_divide_eq)
+by (force simp add: nonzero_divide_eq_eq)
 
 
 subsection {* Ordered Fields *}
--- a/src/HOL/SEQ.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/SEQ.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -40,10 +40,23 @@
 
 definition
   monoseq :: "(nat=>real)=>bool" where
-    --{*Definition for monotonicity*}
+    --{*Definition of monotonicity. 
+        The use of disjunction here complicates proofs considerably. 
+        One alternative is to add a Boolean argument to indicate the direction. 
+        Another is to develop the notions of increasing and decreasing first.*}
   [code del]: "monoseq X = ((\<forall>m. \<forall>n\<ge>m. X m \<le> X n) | (\<forall>m. \<forall>n\<ge>m. X n \<le> X m))"
 
 definition
+  incseq :: "(nat=>real)=>bool" where
+    --{*Increasing sequence*}
+  [code del]: "incseq X = (\<forall>m. \<forall>n\<ge>m. X m \<le> X n)"
+
+definition
+  decseq :: "(nat=>real)=>bool" where
+    --{*Increasing sequence*}
+  [code del]: "decseq X = (\<forall>m. \<forall>n\<ge>m. X n \<le> X m)"
+
+definition
   subseq :: "(nat => nat) => bool" where
     --{*Definition of subsequence*}
   [code del]:   "subseq f = (\<forall>m. \<forall>n>m. (f m) < (f n))"
@@ -886,6 +899,14 @@
   thus ?case by arith
 qed
 
+lemma LIMSEQ_subseq_LIMSEQ:
+  "\<lbrakk> X ----> L; subseq f \<rbrakk> \<Longrightarrow> (X o f) ----> L"
+apply (auto simp add: LIMSEQ_def) 
+apply (drule_tac x=r in spec, clarify)  
+apply (rule_tac x=no in exI, clarify) 
+apply (blast intro: seq_suble le_trans dest!: spec) 
+done
+
 subsection {* Bounded Monotonic Sequences *}
 
 
@@ -1021,6 +1042,47 @@
 apply (auto intro!: Bseq_mono_convergent)
 done
 
+subsubsection{*Increasing and Decreasing Series*}
+
+lemma incseq_imp_monoseq:  "incseq X \<Longrightarrow> monoseq X"
+  by (simp add: incseq_def monoseq_def) 
+
+lemma incseq_le: assumes inc: "incseq X" and lim: "X ----> L" shows "X n \<le> L"
+  using monoseq_le [OF incseq_imp_monoseq [OF inc] lim]
+proof
+  assume "(\<forall>n. X n \<le> L) \<and> (\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n)"
+  thus ?thesis by simp
+next
+  assume "(\<forall>n. L \<le> X n) \<and> (\<forall>m n. m \<le> n \<longrightarrow> X n \<le> X m)"
+  hence const: "(!!m n. m \<le> n \<Longrightarrow> X n = X m)" using inc
+    by (auto simp add: incseq_def intro: order_antisym)
+  have X: "!!n. X n = X 0"
+    by (blast intro: const [of 0]) 
+  have "X = (\<lambda>n. X 0)"
+    by (blast intro: ext X)
+  hence "L = X 0" using LIMSEQ_const [of "X 0"]
+    by (auto intro: LIMSEQ_unique lim) 
+  thus ?thesis
+    by (blast intro: eq_refl X)
+qed
+
+lemma decseq_imp_monoseq:  "decseq X \<Longrightarrow> monoseq X"
+  by (simp add: decseq_def monoseq_def)
+
+lemma decseq_eq_incseq: "decseq X = incseq (\<lambda>n. - X n)" 
+  by (simp add: decseq_def incseq_def)
+
+
+lemma decseq_le: assumes dec: "decseq X" and lim: "X ----> L" shows "L \<le> X n"
+proof -
+  have inc: "incseq (\<lambda>n. - X n)" using dec
+    by (simp add: decseq_eq_incseq)
+  have "- X n \<le> - L" 
+    by (blast intro: incseq_le [OF inc] LIMSEQ_minus lim) 
+  thus ?thesis
+    by simp
+qed
+
 subsubsection{*A Few More Equivalence Theorems for Boundedness*}
 
 text{*alternative formulation for boundedness*}
@@ -1065,6 +1127,14 @@
   "\<lbrakk>Cauchy X; 0 < e\<rbrakk> \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e"
 by (simp add: Cauchy_def)
 
+lemma Cauchy_subseq_Cauchy:
+  "\<lbrakk> Cauchy X; subseq f \<rbrakk> \<Longrightarrow> Cauchy (X o f)"
+apply (auto simp add: Cauchy_def) 
+apply (drule_tac x=e in spec, clarify)  
+apply (rule_tac x=M in exI, clarify) 
+apply (blast intro: seq_suble le_trans dest!: spec) 
+done
+
 subsubsection {* Cauchy Sequences are Bounded *}
 
 text{*A Cauchy sequence is bounded -- this is the standard
@@ -1238,6 +1308,11 @@
   shows "Cauchy X = convergent X"
 by (fast intro: Cauchy_convergent convergent_Cauchy)
 
+lemma convergent_subseq_convergent:
+  fixes X :: "nat \<Rightarrow> 'a::banach"
+  shows "\<lbrakk> convergent X; subseq f \<rbrakk> \<Longrightarrow> convergent (X o f)"
+  by (simp add: Cauchy_subseq_Cauchy Cauchy_convergent_iff [symmetric]) 
+
 
 subsection {* Power Sequences *}
 
--- a/src/HOL/SET-Protocol/Cardholder_Registration.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/SET-Protocol/Cardholder_Registration.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Auth/SET/Cardholder_Registration
-    ID:         $Id$
     Authors:    Giampaolo Bella, Fabio Massacci, Lawrence C Paulson,
                 Piero Tramontano
 *)
@@ -263,7 +262,7 @@
 	 THEN set_cr.SET_CR5 [of concl: C i KC3 NC3 KC2 CardSecret],
 	 THEN Says_to_Gets,  
 	 THEN set_cr.SET_CR6 [of concl: i C KC2]])
-apply (tactic "PublicSET.basic_possibility_tac")
+apply basic_possibility
 apply (simp_all (no_asm_simp) add: symKeys_neq_imp_neq)
 done
 
--- a/src/HOL/SET-Protocol/MessageSET.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/SET-Protocol/MessageSET.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Auth/SET/MessageSET
-    ID:         $Id$
     Authors:     Giampaolo Bella, Fabio Massacci, Lawrence C Paulson
 *)
 
@@ -846,9 +845,9 @@
 (*Prove base case (subgoal i) and simplify others.  A typical base case
   concerns  Crypt K X \<notin> Key`shrK`bad  and cannot be proved by rewriting
   alone.*)
-fun prove_simple_subgoals_tac i =
-    CLASIMPSET' (fn (cs, ss) => force_tac (cs, ss addsimps [@{thm image_eq_UN}])) i THEN
-    ALLGOALS (SIMPSET' asm_simp_tac)
+fun prove_simple_subgoals_tac (cs, ss) i =
+    force_tac (cs, ss addsimps [@{thm image_eq_UN}]) i THEN
+    ALLGOALS (asm_simp_tac ss)
 
 (*Analysis of Fake cases.  Also works for messages that forward unknown parts,
   but this application is no longer necessary if analz_insert_eq is used.
@@ -873,8 +872,7 @@
 		  (cs addIs [@{thm analz_insertI},
 				   impOfSubs @{thm analz_subset_parts}]) 4 1))
 
-(*The explicit claset and simpset arguments help it work with Isar*)
-fun gen_spy_analz_tac (cs,ss) i =
+fun spy_analz_tac (cs,ss) i =
   DETERM
    (SELECT_GOAL
      (EVERY
@@ -887,8 +885,6 @@
        REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])),
        DEPTH_SOLVE (atomic_spy_analz_tac (cs,ss) 1)]) i)
 
-val spy_analz_tac = CLASIMPSET' gen_spy_analz_tac;
-
 end
 *}
 (*>*)
@@ -941,7 +937,7 @@
 
 method_setup spy_analz = {*
     Scan.succeed (fn ctxt =>
-        SIMPLE_METHOD' (MessageSET.gen_spy_analz_tac (local_clasimpset_of ctxt))) *}
+        SIMPLE_METHOD' (MessageSET.spy_analz_tac (local_clasimpset_of ctxt))) *}
     "for proving the Fake case when analz is involved"
 
 method_setup atomic_spy_analz = {*
--- a/src/HOL/SET-Protocol/PublicSET.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/SET-Protocol/PublicSET.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,6 +1,5 @@
 (*  Title:      HOL/Auth/SET/PublicSET
-    ID:         $Id$
-    Authors:     Giampaolo Bella, Fabio Massacci, Lawrence C Paulson
+    Authors:    Giampaolo Bella, Fabio Massacci, Lawrence C Paulson
 *)
 
 header{*The Public-Key Theory, Modified for SET*}
@@ -348,22 +347,19 @@
 structure PublicSET =
 struct
 
-(*Tactic for possibility theorems (Isar interface)*)
-fun gen_possibility_tac ss state = state |>
+(*Tactic for possibility theorems*)
+fun possibility_tac ctxt =
     REPEAT (*omit used_Says so that Nonces start from different traces!*)
-    (ALLGOALS (simp_tac (ss delsimps [@{thm used_Says}, @{thm used_Notes}]))
+    (ALLGOALS (simp_tac (local_simpset_of ctxt delsimps [@{thm used_Says}, @{thm used_Notes}]))
      THEN
      REPEAT_FIRST (eq_assume_tac ORELSE' 
                    resolve_tac [refl, conjI, @{thm Nonce_supply}]))
 
-(*Tactic for possibility theorems (ML script version)*)
-fun possibility_tac state = gen_possibility_tac (simpset_of (Thm.theory_of_thm state)) state
-
 (*For harder protocols (such as SET_CR!), where we have to set up some
   nonces and keys initially*)
-fun basic_possibility_tac st = st |>
+fun basic_possibility_tac ctxt =
     REPEAT 
-    (ALLGOALS (asm_simp_tac (simpset_of (Thm.theory_of_thm st) setSolver safe_solver))
+    (ALLGOALS (asm_simp_tac (local_simpset_of ctxt setSolver safe_solver))
      THEN
      REPEAT_FIRST (resolve_tac [refl, conjI]))
 
@@ -371,10 +367,12 @@
 *}
 
 method_setup possibility = {*
-    Scan.succeed (fn ctxt =>
-        SIMPLE_METHOD (PublicSET.gen_possibility_tac (local_simpset_of ctxt))) *}
+    Scan.succeed (SIMPLE_METHOD o PublicSET.possibility_tac) *}
     "for proving possibility theorems"
 
+method_setup basic_possibility = {*
+    Scan.succeed (SIMPLE_METHOD o PublicSET.basic_possibility_tac) *}
+    "for proving possibility theorems"
 
 
 subsection{*Specialized Rewriting for Theorems About @{term analz} and Image*}
--- a/src/HOL/SET-Protocol/Purchase.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/SET-Protocol/Purchase.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Auth/SET/Purchase
-    ID:         $Id$
     Authors:    Giampaolo Bella, Fabio Massacci, Lawrence C Paulson
 *)
 
@@ -335,7 +334,7 @@
 	  THEN set_pur.AuthResUns [of concl: "PG j" M KP LID_M XID],
           THEN Says_to_Gets, 
 	  THEN set_pur.PRes]) 
-apply (tactic "PublicSET.basic_possibility_tac")
+apply basic_possibility
 apply (simp_all add: used_Cons symKeys_neq_imp_neq) 
 done
 
@@ -371,7 +370,7 @@
 	  THEN set_pur.AuthResS [of concl: "PG j" M KP LID_M XID],
           THEN Says_to_Gets, 
 	  THEN set_pur.PRes]) 
-apply (tactic "PublicSET.basic_possibility_tac")
+apply basic_possibility
 apply (auto simp add: used_Cons symKeys_neq_imp_neq) 
 done
 
--- a/src/HOL/Series.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Series.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -557,7 +557,6 @@
 apply (induct_tac "na", auto)
 apply (rule_tac y = "c * norm (f (N + n))" in order_trans)
 apply (auto intro: mult_right_mono simp add: summable_def)
-apply (simp add: mult_ac)
 apply (rule_tac x = "norm (f N) * (1/ (1 - c)) / (c ^ N)" in exI)
 apply (rule sums_divide) 
 apply (rule sums_mult)
--- a/src/HOL/Set.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Set.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -561,19 +561,15 @@
   "'a set"}.
 *}
 
-lemma subsetD [elim]: "A \<subseteq> B ==> c \<in> A ==> c \<in> B"
+lemma subsetD [elim, intro?]: "A \<subseteq> B ==> c \<in> A ==> c \<in> B"
   -- {* Rule in Modus Ponens style. *}
   by (unfold mem_def) blast
 
-declare subsetD [intro?] -- FIXME
-
-lemma rev_subsetD: "c \<in> A ==> A \<subseteq> B ==> c \<in> B"
+lemma rev_subsetD [intro?]: "c \<in> A ==> A \<subseteq> B ==> c \<in> B"
   -- {* The same, with reversed premises for use with @{text erule} --
       cf @{text rev_mp}. *}
   by (rule subsetD)
 
-declare rev_subsetD [intro?] -- FIXME
-
 text {*
   \medskip Converts @{prop "A \<subseteq> B"} to @{prop "x \<in> A ==> x \<in> B"}.
 *}
@@ -623,8 +619,6 @@
   -- {* Anti-symmetry of the subset relation. *}
   by (iprover intro: set_ext subsetD)
 
-lemmas equalityI [intro!] = subset_antisym
-
 text {*
   \medskip Equality rules from ZF set theory -- are they appropriate
   here?
@@ -1064,11 +1058,6 @@
 
 lemmas split_ifs = if_bool_eq_conj split_if_eq1 split_if_eq2 split_if_mem1 split_if_mem2
 
-lemmas mem_simps =
-  insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff
-  mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff
-  -- {* Each of these has ALREADY been added @{text "[simp]"} above. *}
-
 (*Would like to add these, but the existing code only searches for the
   outer-level constant, which in this case is just "op :"; we instead need
   to use term-nets to associate patterns with rules.  Also, if a rule fails to
@@ -2489,7 +2478,13 @@
   Sup  ("\<Squnion>_" [900] 900)
 
 
-subsection {* Basic ML bindings *}
+subsection {* Misc theorem and ML bindings *}
+
+lemmas equalityI = subset_antisym
+lemmas mem_simps =
+  insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff
+  mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff
+  -- {* Each of these has ALREADY been added @{text "[simp]"} above. *}
 
 ML {*
 val Ball_def = @{thm Ball_def}
--- a/src/HOL/SizeChange/sct.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/SizeChange/sct.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/SizeChange/sct.ML
-    ID:         $Id$
     Author:     Alexander Krauss, TU Muenchen
 
 Tactics for size change termination.
@@ -106,7 +105,7 @@
 
 fun dest_case rebind t =
     let
-      val (_ $ _ $ rhs :: _ $ _ $ match :: guards) = HOLogic.dest_conj t
+      val ((_ $ _ $ rhs) :: (_ $ _ $ match) :: guards) = HOLogic.dest_conj t
       val guard = case guards of [] => HOLogic.true_const | gs => foldr1 HOLogic.mk_conj gs
     in
       foldr1 HOLogic.mk_prod [rebind guard, rebind rhs, rebind match]
@@ -183,8 +182,10 @@
 
 fun rand (_ $ t) = t
 
-fun setup_probe_goal thy domT Dtab Mtab (i, j) =
+fun setup_probe_goal ctxt domT Dtab Mtab (i, j) =
     let
+      val css = local_clasimpset_of ctxt
+      val thy = ProofContext.theory_of ctxt
       val RD1 = get_elem (Dtab i)
       val RD2 = get_elem (Dtab j)
       val Ms1 = get_elem (Mtab i)
@@ -200,12 +201,12 @@
       val saved_state = HOLogic.mk_Trueprop (mk_stepP RD1 RD2 mvar1 mvar2 relvar)
                          |> cterm_of thy
                          |> Goal.init
-                         |> CLASIMPSET auto_tac |> Seq.hd
+                         |> auto_tac css |> Seq.hd
 
       val no_step = saved_state
                       |> forall_intr (cterm_of thy relvar)
                       |> forall_elim (cterm_of thy (Abs ("", HOLogic.natT, Abs ("", HOLogic.natT, HOLogic.false_const))))
-                      |> CLASIMPSET auto_tac |> Seq.hd
+                      |> auto_tac css |> Seq.hd
 
     in
       if Thm.no_prems no_step
@@ -218,7 +219,7 @@
                 val with_m1 = saved_state
                                 |> forall_intr (cterm_of thy mvar1)
                                 |> forall_elim (cterm_of thy M1)
-                                |> CLASIMPSET auto_tac |> Seq.hd
+                                |> auto_tac css |> Seq.hd
 
                 fun set_m2 j =
                     let
@@ -226,15 +227,15 @@
                       val with_m2 = with_m1
                                       |> forall_intr (cterm_of thy mvar2)
                                       |> forall_elim (cterm_of thy M2)
-                                      |> CLASIMPSET auto_tac |> Seq.hd
+                                      |> auto_tac css |> Seq.hd
 
                       val decr = forall_intr (cterm_of thy relvar)
                                    #> forall_elim (cterm_of thy @{const HOL.less(nat)})
-                                   #> CLASIMPSET auto_tac #> Seq.hd
+                                   #> auto_tac css #> Seq.hd
 
                       val decreq = forall_intr (cterm_of thy relvar)
                                      #> forall_elim (cterm_of thy @{const HOL.less_eq(nat)})
-                                     #> CLASIMPSET auto_tac #> Seq.hd
+                                     #> auto_tac css #> Seq.hd
 
                       val thm1 = decr with_m2
                     in
@@ -266,17 +267,17 @@
 
 fun mk_edge m G n = HOLogic.mk_prod (m, HOLogic.mk_prod (G, n))
 
-val in_graph_tac =
+fun in_graph_tac ctxt =
     simp_tac (HOL_basic_ss addsimps has_edge_simps) 1
-    THEN SIMPSET (fn x => simp_tac x 1) (* FIXME reduce simpset *)
+    THEN (simp_tac (local_simpset_of ctxt) 1) (* FIXME reduce simpset *)
 
-fun approx_tac (NoStep thm) = rtac disjI1 1 THEN rtac thm 1
-  | approx_tac (Graph (G, thm)) =
+fun approx_tac _ (NoStep thm) = rtac disjI1 1 THEN rtac thm 1
+  | approx_tac ctxt (Graph (G, thm)) =
     rtac disjI2 1
     THEN rtac exI 1
     THEN rtac conjI 1
     THEN rtac thm 2
-    THEN in_graph_tac
+    THEN (in_graph_tac ctxt)
 
 fun all_less_tac [] = rtac all_less_zero 1
   | all_less_tac (t :: ts) = rtac all_less_Suc 1
@@ -292,7 +293,7 @@
 
 fun mk_call_graph ctxt (st : thm) =
     let
-      val thy = theory_of_thm st
+      val thy = ProofContext.theory_of ctxt
       val _ $ _ $ RDlist $ _ = HOLogic.dest_Trueprop (hd (prems_of st))
 
       val RDs = HOLogic.dest_list RDlist
@@ -316,7 +317,7 @@
       val pairs = matrix indices indices
       val parts = map_matrix (fn (n,m) =>
                                  (timeap_msg (string_of_int n ^ "," ^ string_of_int m)
-                                             (setup_probe_goal thy domT Dtab Mtab) (n,m))) pairs
+                                             (setup_probe_goal ctxt domT Dtab Mtab) (n,m))) pairs
 
 
       val s = fold_index (fn (i, cs) => fold_index (fn (j, Graph (G, _)) => prefix ("(" ^ string_of_int i ^ "," ^ string_of_int j ^ "): " ^
@@ -333,8 +334,8 @@
       val sound_int_goal = HOLogic.mk_Trueprop (mk_sound_int ACG RDlist mfuns)
 
       val tac =
-          (SIMPSET (unfold_tac [sound_int_def, len_simp]))
-            THEN all_less_tac (map (all_less_tac o map approx_tac) parts)
+          unfold_tac [sound_int_def, len_simp] (local_simpset_of ctxt)
+            THEN all_less_tac (map (all_less_tac o map (approx_tac ctxt)) parts)
     in
       tac (instantiate' [] [SOME (cterm_of thy ACG), SOME (cterm_of thy mfuns)] st)
     end
--- a/src/HOL/Tools/Groebner_Basis/normalizer_data.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/Groebner_Basis/normalizer_data.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -191,16 +191,17 @@
 
 in
 
-val attribute =
-  Scan.lift (Args.$$$ delN >> K del) ||
-    ((keyword2 semiringN opsN |-- terms) --
-     (keyword2 semiringN rulesN |-- thms)) --
-    (optional (keyword2 ringN opsN |-- terms) --
-     optional (keyword2 ringN rulesN |-- thms)) --
-    optional (keyword2 idomN rulesN |-- thms) --
-    optional (keyword2 idealN rulesN |-- thms)
-    >> (fn (((sr, r), id), idl) => 
-          add {semiring = sr, ring = r, idom = id, ideal = idl});
+val normalizer_setup =
+  Attrib.setup @{binding normalizer}
+    (Scan.lift (Args.$$$ delN >> K del) ||
+      ((keyword2 semiringN opsN |-- terms) --
+       (keyword2 semiringN rulesN |-- thms)) --
+      (optional (keyword2 ringN opsN |-- terms) --
+       optional (keyword2 ringN rulesN |-- thms)) --
+      optional (keyword2 idomN rulesN |-- thms) --
+      optional (keyword2 idealN rulesN |-- thms)
+      >> (fn (((sr, r), id), idl) => add {semiring = sr, ring = r, idom = id, ideal = idl}))
+    "semiring normalizer data";
 
 end;
 
@@ -208,7 +209,7 @@
 (* theory setup *)
 
 val setup =
-  Attrib.setup @{binding normalizer} attribute "semiring normalizer data" #>
+  normalizer_setup #>
   Attrib.setup @{binding algebra} (Attrib.add_del add_ss del_ss) "pre-simplification for algebra";
 
 end;
--- a/src/HOL/Tools/Qelim/cooper.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/Qelim/cooper.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -23,7 +23,7 @@
 val false_tm = @{cterm "False"};
 val zdvd1_eq = @{thm "zdvd1_eq"};
 val presburger_ss = @{simpset} addsimps [zdvd1_eq];
-val lin_ss = presburger_ss addsimps (@{thm "dvd_eq_mod_eq_0"}::zdvd1_eq::@{thms zadd_ac});
+val lin_ss = presburger_ss addsimps (@{thm dvd_eq_mod_eq_0} :: zdvd1_eq :: @{thms zadd_ac});
 
 val iT = HOLogic.intT
 val bT = HOLogic.boolT;
@@ -172,7 +172,7 @@
 
     (* Canonical linear form for terms, formulae etc.. *)
 fun provelin ctxt t = Goal.prove ctxt [] [] t 
-  (fn _ => EVERY [simp_tac lin_ss 1, TRY (simple_arith_tac ctxt 1)]);
+  (fn _ => EVERY [simp_tac lin_ss 1, TRY (linear_arith_tac ctxt 1)]);
 fun linear_cmul 0 tm = zero 
   | linear_cmul n tm = case tm of  
       Const (@{const_name HOL.plus}, _) $ a $ b => addC $ linear_cmul n a $ linear_cmul n b
--- a/src/HOL/Tools/Qelim/cooper_data.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/Qelim/cooper_data.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -67,6 +67,7 @@
 (* concrete syntax *)
 
 local
+
 fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
 
 val constsN = "consts";
@@ -77,14 +78,17 @@
 fun optional scan = Scan.optional scan [];
 
 in
-val attribute =
- (Scan.lift (Args.$$$ "del") |-- optional (keyword constsN |-- terms)) >> del ||
-  optional (keyword constsN |-- terms) >> add
+
+val presburger_setup =
+  Attrib.setup @{binding presburger}
+    ((Scan.lift (Args.$$$ "del") |-- optional (keyword constsN |-- terms)) >> del ||
+      optional (keyword constsN |-- terms) >> add) "Cooper data";
+
 end;
 
 
 (* theory setup *)
 
-val setup = Attrib.setup @{binding presburger} attribute "Cooper data";
+val setup = presburger_setup;
 
 end;
--- a/src/HOL/Tools/Qelim/ferrante_rackoff_data.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/Qelim/ferrante_rackoff_data.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -122,26 +122,28 @@
 val terms = thms >> map Drule.dest_term;
 in
 
-val attribute =
-    (keyword minfN |-- thms)
+val ferrak_setup =
+  Attrib.setup @{binding ferrack}
+    ((keyword minfN |-- thms)
      -- (keyword pinfN |-- thms)
      -- (keyword nmiN |-- thms)
      -- (keyword npiN |-- thms)
      -- (keyword lin_denseN |-- thms)
      -- (keyword qeN |-- thms)
-     -- (keyword atomsN |-- terms) >> 
-     (fn ((((((minf,pinf),nmi),npi),lin_dense),qe), atoms)=> 
-     if length qe = 1 then
-       add ({minf = minf, pinf = pinf, nmi = nmi, npi = npi, ld = lin_dense, 
-            qe = hd qe, atoms = atoms},
-           {isolate_conv = undefined, whatis = undefined, simpset = HOL_ss})
-     else error "only one theorem for qe!")
+     -- (keyword atomsN |-- terms) >>
+       (fn ((((((minf,pinf),nmi),npi),lin_dense),qe), atoms)=> 
+       if length qe = 1 then
+         add ({minf = minf, pinf = pinf, nmi = nmi, npi = npi, ld = lin_dense, 
+              qe = hd qe, atoms = atoms},
+             {isolate_conv = undefined, whatis = undefined, simpset = HOL_ss})
+       else error "only one theorem for qe!"))
+    "Ferrante Rackoff data";
 
 end;
 
 
 (* theory setup *)
 
-val setup = Attrib.setup @{binding ferrack} attribute "Ferrante Rackoff data";
+val setup = ferrak_setup;
 
 end;
--- a/src/HOL/Tools/Qelim/langford_data.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/Qelim/langford_data.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -85,25 +85,28 @@
 val terms = thms >> map Drule.dest_term;
 in
 
-val attribute =
-    (keyword qeN |-- thms)
+val langford_setup =
+  Attrib.setup @{binding langford}
+    ((keyword qeN |-- thms)
      -- (keyword gatherN |-- thms)
-     -- (keyword atomsN |-- terms) >> 
-     (fn ((qes,gs), atoms) => 
-     if length qes = 3 andalso length gs > 1 then
-       let 
-         val [q1,q2,q3] = qes
-         val gst::gss = gs
-         val entry = {qe_bnds = q1, qe_nolb = q2,
-                      qe_noub = q3, gs = gss, gst = gst, atoms = atoms}
-       in add entry end
-     else error "Need 3 theorems for qe and at least one for gs")
+     -- (keyword atomsN |-- terms) >>
+       (fn ((qes,gs), atoms) => 
+       if length qes = 3 andalso length gs > 1 then
+         let 
+           val [q1,q2,q3] = qes
+           val gst::gss = gs
+           val entry = {qe_bnds = q1, qe_nolb = q2,
+                        qe_noub = q3, gs = gss, gst = gst, atoms = atoms}
+         in add entry end
+       else error "Need 3 theorems for qe and at least one for gs"))
+    "Langford data";
+
 end;
 
 (* theory setup *)
 
 val setup =
-  Attrib.setup @{binding langford} attribute "Langford data" #>
+  langford_setup #>
   Attrib.setup @{binding langfordsimp} (Attrib.add_del add_simp del_simp) "Langford simpset";
 
 end;
--- a/src/HOL/Tools/Qelim/presburger.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/Qelim/presburger.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -163,8 +163,10 @@
 
 fun cooper_tac elim add_ths del_ths ctxt =
 let val ss = Simplifier.context ctxt (fst (CooperData.get ctxt)) delsimps del_ths addsimps add_ths
+    val aprems = Arith_Data.get_arith_facts ctxt
 in
-  ObjectLogic.full_atomize_tac
+  Method.insert_tac aprems
+  THEN_ALL_NEW ObjectLogic.full_atomize_tac
   THEN_ALL_NEW CONVERSION Thm.eta_long_conversion
   THEN_ALL_NEW simp_tac ss
   THEN_ALL_NEW (TRY o generalize_tac (int_nat_terms ctxt))
--- a/src/HOL/Tools/Qelim/qelim.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/Qelim/qelim.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/Qelim/qelim.ML
-    ID:         $Id$
     Author:     Amine Chaieb, TU Muenchen
 
 Generic quantifier elimination conversions for HOL formulae.
@@ -26,11 +25,12 @@
    case (term_of p) of
     Const(s,T)$_$_ => 
        if domain_type T = HOLogic.boolT
-          andalso s mem ["op &","op |","op -->","op ="]
+          andalso member (op =) [@{const_name "op &"}, @{const_name "op |"},
+            @{const_name "op -->"}, @{const_name "op ="}] s
        then binop_conv (conv env) p 
        else atcv env p
-  | Const("Not",_)$_ => arg_conv (conv env) p
-  | Const("Ex",_)$Abs(s,_,_) =>
+  | Const(@{const_name "Not"},_)$_ => arg_conv (conv env) p
+  | Const(@{const_name "Ex"},_)$Abs(s,_,_) =>
     let
      val (e,p0) = Thm.dest_comb p
      val (x,p') = Thm.dest_abs (SOME s) p0
@@ -41,8 +41,8 @@
      val (l,r) = Thm.dest_equals (cprop_of th')
     in if Thm.is_reflexive th' then Thm.transitive th (qcv env (Thm.rhs_of th))
        else Thm.transitive (Thm.transitive th th') (conv env r) end
-  | Const("Ex",_)$ _ => (Thm.eta_long_conversion then_conv conv env) p
-  | Const("All",_)$_ =>
+  | Const(@{const_name "Ex"},_)$ _ => (Thm.eta_long_conversion then_conv conv env) p
+  | Const(@{const_name "All"},_)$_ =>
     let
      val p = Thm.dest_arg p
      val ([(_,T)],[]) = Thm.match (@{cpat "All"}, Thm.dest_fun p)
--- a/src/HOL/Tools/TFL/post.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/TFL/post.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -55,7 +55,7 @@
   Prim.postprocess strict
    {wf_tac     = REPEAT (ares_tac wfs 1),
     terminator = asm_simp_tac ss 1
-                 THEN TRY (silent_arith_tac (Simplifier.the_context ss) 1 ORELSE
+                 THEN TRY (Arith_Data.arith_tac (Simplifier.the_context ss) 1 ORELSE
                            fast_tac (cs addSDs [@{thm not0_implies_Suc}] addss ss) 1),
     simplifier = Rules.simpl_conv ss []};
 
@@ -79,10 +79,7 @@
    in lhs aconv rhs end
    handle U.ERR _ => false;
    
-
-fun prover s = prove_goal @{theory HOL} s (fn _ => [fast_tac HOL_cs 1]);
-val P_imp_P_iff_True = prover "P --> (P= True)" RS mp;
-val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection;
+val P_imp_P_eq_True = @{thm eqTrueI} RS eq_reflection;
 fun mk_meta_eq r = case concl_of r of
      Const("==",_)$_$_ => r
   |   _ $(Const("op =",_)$_$_) => r RS eq_reflection
--- a/src/HOL/Tools/arith_data.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/arith_data.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -6,6 +6,11 @@
 
 signature ARITH_DATA =
 sig
+  val arith_tac: Proof.context -> int -> tactic
+  val verbose_arith_tac: Proof.context -> int -> tactic
+  val add_tactic: string -> (bool -> Proof.context -> int -> tactic) -> theory -> theory
+  val get_arith_facts: Proof.context -> thm list
+
   val prove_conv_nohyps: tactic list -> Proof.context -> term * term -> thm option
   val prove_conv: tactic list -> Proof.context -> thm list -> term * term -> thm option
   val prove_conv2: tactic -> (simpset -> tactic) -> simpset -> term * term -> thm
@@ -14,11 +19,55 @@
   val trans_tac: thm option -> tactic
   val prep_simproc: string * string list * (theory -> simpset -> term -> thm option)
     -> simproc
+
+  val setup: theory -> theory
 end;
 
 structure Arith_Data: ARITH_DATA =
 struct
 
+(* slots for pluging in arithmetic facts and tactics *)
+
+structure Arith_Facts = NamedThmsFun(
+  val name = "arith"
+  val description = "arith facts - only ground formulas"
+);
+
+val get_arith_facts = Arith_Facts.get;
+
+structure Arith_Tactics = TheoryDataFun
+(
+  type T = (serial * (string * (bool -> Proof.context -> int -> tactic))) list;
+  val empty = [];
+  val copy = I;
+  val extend = I;
+  fun merge _ = AList.merge (op =) (K true);
+);
+
+fun add_tactic name tac = Arith_Tactics.map (cons (serial (), (name, tac)));
+
+fun gen_arith_tac verbose ctxt =
+  let
+    val tactics = (Arith_Tactics.get o ProofContext.theory_of) ctxt
+    fun invoke (_, (name, tac)) k st = (if verbose
+      then warning ("Trying " ^ name ^ "...") else ();
+      tac verbose ctxt k st);
+  in FIRST' (map invoke (rev tactics)) end;
+
+val arith_tac = gen_arith_tac false;
+val verbose_arith_tac = gen_arith_tac true;
+
+val setup =
+  Arith_Facts.setup #>
+  Method.setup @{binding arith}
+    (Args.bang_facts >> (fn prems => fn ctxt =>
+      METHOD (fn facts => HEADGOAL (Method.insert_tac (prems @ get_arith_facts ctxt @ facts)
+        THEN' verbose_arith_tac ctxt))))
+    "various arithmetic decision procedures";
+
+
+(* various auxiliary and legacy *)
+
 fun prove_conv_nohyps tacs ctxt (t, u) =
   if t aconv u then NONE
   else let val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u))
--- a/src/HOL/Tools/cnf_funcs.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/cnf_funcs.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -51,51 +51,46 @@
 structure cnf : CNF =
 struct
 
-fun thm_by_auto (G : string) : thm =
-	prove_goal (the_context ()) G (fn prems => [cut_facts_tac prems 1, CLASIMPSET auto_tac]);
+val clause2raw_notE      = @{lemma "[| P; ~P |] ==> False" by auto};
+val clause2raw_not_disj  = @{lemma "[| ~P; ~Q |] ==> ~(P | Q)" by auto};
+val clause2raw_not_not   = @{lemma "P ==> ~~P" by auto};
 
-(* Thm.thm *)
-val clause2raw_notE      = thm_by_auto "[| P; ~P |] ==> False";
-val clause2raw_not_disj  = thm_by_auto "[| ~P; ~Q |] ==> ~(P | Q)";
-val clause2raw_not_not   = thm_by_auto "P ==> ~~P";
+val iff_refl             = @{lemma "(P::bool) = P" by auto};
+val iff_trans            = @{lemma "[| (P::bool) = Q; Q = R |] ==> P = R" by auto};
+val conj_cong            = @{lemma "[| P = P'; Q = Q' |] ==> (P & Q) = (P' & Q')" by auto};
+val disj_cong            = @{lemma "[| P = P'; Q = Q' |] ==> (P | Q) = (P' | Q')" by auto};
 
-val iff_refl             = thm_by_auto "(P::bool) = P";
-val iff_trans            = thm_by_auto "[| (P::bool) = Q; Q = R |] ==> P = R";
-val conj_cong            = thm_by_auto "[| P = P'; Q = Q' |] ==> (P & Q) = (P' & Q')";
-val disj_cong            = thm_by_auto "[| P = P'; Q = Q' |] ==> (P | Q) = (P' | Q')";
-
-val make_nnf_imp         = thm_by_auto "[| (~P) = P'; Q = Q' |] ==> (P --> Q) = (P' | Q')";
-val make_nnf_iff         = thm_by_auto "[| P = P'; (~P) = NP; Q = Q'; (~Q) = NQ |] ==> (P = Q) = ((P' | NQ) & (NP | Q'))";
-val make_nnf_not_false   = thm_by_auto "(~False) = True";
-val make_nnf_not_true    = thm_by_auto "(~True) = False";
-val make_nnf_not_conj    = thm_by_auto "[| (~P) = P'; (~Q) = Q' |] ==> (~(P & Q)) = (P' | Q')";
-val make_nnf_not_disj    = thm_by_auto "[| (~P) = P'; (~Q) = Q' |] ==> (~(P | Q)) = (P' & Q')";
-val make_nnf_not_imp     = thm_by_auto "[| P = P'; (~Q) = Q' |] ==> (~(P --> Q)) = (P' & Q')";
-val make_nnf_not_iff     = thm_by_auto "[| P = P'; (~P) = NP; Q = Q'; (~Q) = NQ |] ==> (~(P = Q)) = ((P' | Q') & (NP | NQ))";
-val make_nnf_not_not     = thm_by_auto "P = P' ==> (~~P) = P'";
+val make_nnf_imp         = @{lemma "[| (~P) = P'; Q = Q' |] ==> (P --> Q) = (P' | Q')" by auto};
+val make_nnf_iff         = @{lemma "[| P = P'; (~P) = NP; Q = Q'; (~Q) = NQ |] ==> (P = Q) = ((P' | NQ) & (NP | Q'))" by auto};
+val make_nnf_not_false   = @{lemma "(~False) = True" by auto};
+val make_nnf_not_true    = @{lemma "(~True) = False" by auto};
+val make_nnf_not_conj    = @{lemma "[| (~P) = P'; (~Q) = Q' |] ==> (~(P & Q)) = (P' | Q')" by auto};
+val make_nnf_not_disj    = @{lemma "[| (~P) = P'; (~Q) = Q' |] ==> (~(P | Q)) = (P' & Q')" by auto};
+val make_nnf_not_imp     = @{lemma "[| P = P'; (~Q) = Q' |] ==> (~(P --> Q)) = (P' & Q')" by auto};
+val make_nnf_not_iff     = @{lemma "[| P = P'; (~P) = NP; Q = Q'; (~Q) = NQ |] ==> (~(P = Q)) = ((P' | Q') & (NP | NQ))" by auto};
+val make_nnf_not_not     = @{lemma "P = P' ==> (~~P) = P'" by auto};
 
-val simp_TF_conj_True_l  = thm_by_auto "[| P = True; Q = Q' |] ==> (P & Q) = Q'";
-val simp_TF_conj_True_r  = thm_by_auto "[| P = P'; Q = True |] ==> (P & Q) = P'";
-val simp_TF_conj_False_l = thm_by_auto "P = False ==> (P & Q) = False";
-val simp_TF_conj_False_r = thm_by_auto "Q = False ==> (P & Q) = False";
-val simp_TF_disj_True_l  = thm_by_auto "P = True ==> (P | Q) = True";
-val simp_TF_disj_True_r  = thm_by_auto "Q = True ==> (P | Q) = True";
-val simp_TF_disj_False_l = thm_by_auto "[| P = False; Q = Q' |] ==> (P | Q) = Q'";
-val simp_TF_disj_False_r = thm_by_auto "[| P = P'; Q = False |] ==> (P | Q) = P'";
+val simp_TF_conj_True_l  = @{lemma "[| P = True; Q = Q' |] ==> (P & Q) = Q'" by auto};
+val simp_TF_conj_True_r  = @{lemma "[| P = P'; Q = True |] ==> (P & Q) = P'" by auto};
+val simp_TF_conj_False_l = @{lemma "P = False ==> (P & Q) = False" by auto};
+val simp_TF_conj_False_r = @{lemma "Q = False ==> (P & Q) = False" by auto};
+val simp_TF_disj_True_l  = @{lemma "P = True ==> (P | Q) = True" by auto};
+val simp_TF_disj_True_r  = @{lemma "Q = True ==> (P | Q) = True" by auto};
+val simp_TF_disj_False_l = @{lemma "[| P = False; Q = Q' |] ==> (P | Q) = Q'" by auto};
+val simp_TF_disj_False_r = @{lemma "[| P = P'; Q = False |] ==> (P | Q) = P'" by auto};
 
-val make_cnf_disj_conj_l = thm_by_auto "[| (P | R) = PR; (Q | R) = QR |] ==> ((P & Q) | R) = (PR & QR)";
-val make_cnf_disj_conj_r = thm_by_auto "[| (P | Q) = PQ; (P | R) = PR |] ==> (P | (Q & R)) = (PQ & PR)";
+val make_cnf_disj_conj_l = @{lemma "[| (P | R) = PR; (Q | R) = QR |] ==> ((P & Q) | R) = (PR & QR)" by auto};
+val make_cnf_disj_conj_r = @{lemma "[| (P | Q) = PQ; (P | R) = PR |] ==> (P | (Q & R)) = (PQ & PR)" by auto};
 
-val make_cnfx_disj_ex_l = thm_by_auto "((EX (x::bool). P x) | Q) = (EX x. P x | Q)";
-val make_cnfx_disj_ex_r = thm_by_auto "(P | (EX (x::bool). Q x)) = (EX x. P | Q x)";
-val make_cnfx_newlit    = thm_by_auto "(P | Q) = (EX x. (P | x) & (Q | ~x))";
-val make_cnfx_ex_cong   = thm_by_auto "(ALL (x::bool). P x = Q x) ==> (EX x. P x) = (EX x. Q x)";
+val make_cnfx_disj_ex_l  = @{lemma "((EX (x::bool). P x) | Q) = (EX x. P x | Q)" by auto};
+val make_cnfx_disj_ex_r  = @{lemma "(P | (EX (x::bool). Q x)) = (EX x. P | Q x)" by auto};
+val make_cnfx_newlit     = @{lemma "(P | Q) = (EX x. (P | x) & (Q | ~x))" by auto};
+val make_cnfx_ex_cong    = @{lemma "(ALL (x::bool). P x = Q x) ==> (EX x. P x) = (EX x. Q x)" by auto};
 
-val weakening_thm        = thm_by_auto "[| P; Q |] ==> Q";
+val weakening_thm        = @{lemma "[| P; Q |] ==> Q" by auto};
 
-val cnftac_eq_imp        = thm_by_auto "[| P = Q; P |] ==> Q";
+val cnftac_eq_imp        = @{lemma "[| P = Q; P |] ==> Q" by auto};
 
-(* Term.term -> bool *)
 fun is_atom (Const ("False", _))                                           = false
   | is_atom (Const ("True", _))                                            = false
   | is_atom (Const ("op &", _) $ _ $ _)                                    = false
@@ -105,11 +100,9 @@
   | is_atom (Const ("Not", _) $ _)                                         = false
   | is_atom _                                                              = true;
 
-(* Term.term -> bool *)
 fun is_literal (Const ("Not", _) $ x) = is_atom x
   | is_literal x                      = is_atom x;
 
-(* Term.term -> bool *)
 fun is_clause (Const ("op |", _) $ x $ y) = is_clause x andalso is_clause y
   | is_clause x                           = is_literal x;
 
--- a/src/HOL/Tools/function_package/fundef_core.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/function_package/fundef_core.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -691,8 +691,9 @@
 
 
 (* FIXME: This should probably use fixed goals, to be more reliable and faster *)
-fun mk_domain_intro thy (Globals {domT, ...}) R R_cases clause =
+fun mk_domain_intro ctxt (Globals {domT, ...}) R R_cases clause =
     let
+      val thy = ProofContext.theory_of ctxt
       val ClauseInfo {cdata = ClauseContext {qs, gs, lhs, rhs, cqs, ...},
                       qglr = (oqs, _, _, _), ...} = clause
       val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs)
@@ -702,7 +703,7 @@
       Goal.init goal
       |> (SINGLE (resolve_tac [accI] 1)) |> the
       |> (SINGLE (eresolve_tac [Thm.forall_elim_vars 0 R_cases] 1))  |> the
-      |> (SINGLE (CLASIMPSET auto_tac)) |> the
+      |> (SINGLE (auto_tac (local_clasimpset_of ctxt))) |> the
       |> Goal.conclude
       |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
     end
@@ -831,7 +832,7 @@
                          ((rtac (G_induct OF [a]))
                             THEN_ALL_NEW (rtac accI)
                             THEN_ALL_NEW (etac R_cases)
-                            THEN_ALL_NEW (SIMPSET' asm_full_simp_tac)) 1)
+                            THEN_ALL_NEW (asm_full_simp_tac (local_simpset_of octxt))) 1)
 
       val default_thm = (forall_intr_vars graph_implies_dom) COMP (f_def COMP fundef_default_value)
 
@@ -849,9 +850,9 @@
                        (fn _ =>
                            rtac (instantiate' [] [SOME (cterm_of thy lhs_acc)] case_split) 1
                                 THEN (rtac (Thm.forall_elim_vars 0 psimp) THEN_ALL_NEW assume_tac) 1
-                                THEN (SIMPSET' simp_default_tac 1)
+                                THEN (simp_default_tac (local_simpset_of ctxt) 1)
                                 THEN (etac not_acc_down 1)
-                                THEN ((etac R_cases) THEN_ALL_NEW (SIMPSET' simp_default_tac)) 1)
+                                THEN ((etac R_cases) THEN_ALL_NEW (simp_default_tac (local_simpset_of ctxt))) 1)
               |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
           end
     in
@@ -935,7 +936,7 @@
             val total_intro = PROFILE "Proving nested termination rule" (mk_nest_term_rule newthy globals R R_elim) xclauses
 
             val dom_intros = if domintros
-                             then SOME (PROFILE "Proving domain introduction rules" (map (mk_domain_intro newthy globals R R_elim)) xclauses)
+                             then SOME (PROFILE "Proving domain introduction rules" (map (mk_domain_intro lthy globals R R_elim)) xclauses)
                              else NONE
             val trsimps = if tailrec then SOME (mk_trsimps psimps) else NONE
 
--- a/src/HOL/Tools/function_package/lexicographic_order.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/function_package/lexicographic_order.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -176,7 +176,7 @@
 
       val gc = map (fn i => chr (i + 96)) (1 upto length table)
       val mc = 1 upto length measure_funs
-      val tstr = "Result matrix:" ::  "   " ^ concat (map (enclose " " " " o string_of_int) mc)
+      val tstr = "Result matrix:" ::  ("   " ^ concat (map (enclose " " " " o string_of_int) mc))
                  :: map2 (fn r => fn i => i ^ ": " ^ concat (map pr_cell r)) table gc
       val gstr = "Calls:" :: map2 (prefix "  " oo pr_goal) tl gc
       val mstr = "Measures:" :: map2 (prefix "  " oo pr_fun) measure_funs mc
--- a/src/HOL/Tools/function_package/mutual.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/function_package/mutual.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/function_package/mutual.ML
-    ID:         $Id$
     Author:     Alexander Krauss, TU Muenchen
 
 A package for general recursive function definitions.
@@ -207,7 +206,7 @@
                  (HOLogic.Trueprop $ HOLogic.mk_eq (list_comb (f, args), rhs))
                  (fn _ => (LocalDefs.unfold_tac ctxt all_orig_fdefs)
                           THEN EqSubst.eqsubst_tac ctxt [0] [simp] 1
-                          THEN SIMPSET' (fn ss => simp_tac (ss addsimps SumTree.proj_in_rules)) 1)
+                          THEN (simp_tac (local_simpset_of ctxt addsimps SumTree.proj_in_rules)) 1)
         |> restore_cond 
         |> export
     end
--- a/src/HOL/Tools/function_package/scnp_reconstruct.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/function_package/scnp_reconstruct.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -197,7 +197,7 @@
               else if b <= a then @{thm pair_leqI2} else @{thm pair_leqI1}
           in
             rtac rule 1 THEN PRIMITIVE (Thm.elim_implies stored_thm)
-            THEN (if tag_flag then arith_tac ctxt 1 else all_tac)
+            THEN (if tag_flag then Arith_Data.verbose_arith_tac ctxt 1 else all_tac)
           end
 
         fun steps_tac MAX strict lq lp =
@@ -291,7 +291,7 @@
          THEN (rtac @{thm rp_inv_image_rp} 1)
          THEN (rtac (order_rpair ms_rp label) 1)
          THEN PRIMITIVE (instantiate' [] [SOME level_mapping])
-         THEN unfold_tac @{thms rp_inv_image_def} (simpset_of thy)
+         THEN unfold_tac @{thms rp_inv_image_def} (local_simpset_of ctxt)
          THEN LocalDefs.unfold_tac ctxt
            (@{thms split_conv} @ @{thms fst_conv} @ @{thms snd_conv})
          THEN REPEAT (SOMEGOAL (resolve_tac [@{thm Un_least}, @{thm empty_subsetI}]))
@@ -395,7 +395,7 @@
 
 fun gen_sizechange_tac orders autom_tac ctxt err_cont =
   TRY (FundefCommon.apply_termination_rule ctxt 1)
-  THEN TRY Termination.wf_union_tac
+  THEN TRY (Termination.wf_union_tac ctxt)
   THEN
    (rtac @{thm wf_empty} 1
     ORELSE gen_decomp_scnp_tac orders autom_tac ctxt err_cont 1)
--- a/src/HOL/Tools/function_package/termination.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/function_package/termination.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -40,7 +40,7 @@
 
   val REPEAT : ttac -> ttac
 
-  val wf_union_tac : tactic
+  val wf_union_tac : Proof.context -> tactic
 end
 
 
@@ -276,9 +276,9 @@
 
 in
 
-fun wf_union_tac st =
+fun wf_union_tac ctxt st =
     let
-      val thy = theory_of_thm st
+      val thy = ProofContext.theory_of ctxt
       val cert = cterm_of (theory_of_thm st)
       val ((trueprop $ (wf $ rel)) :: ineqs) = prems_of st
 
@@ -303,7 +303,7 @@
           THEN' ((rtac @{thm refl})                       (* unification instantiates all Vars *)
                  ORELSE' ((rtac @{thm conjI})
                           THEN' (rtac @{thm refl})
-                          THEN' (CLASET' blast_tac)))     (* Solve rest of context... not very elegant *)
+                          THEN' (blast_tac (local_claset_of ctxt))))  (* Solve rest of context... not very elegant *)
           ) i
     in
       ((PRIMITIVE (Drule.cterm_instantiate [(cert rel, cert relation)])
--- a/src/HOL/Tools/inductive_package.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/inductive_package.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -460,17 +460,18 @@
 val inductive_cases_i = gen_inductive_cases (K I) Syntax.check_prop;
 
 
-val ind_cases =
-  Scan.lift (Scan.repeat1 Args.name_source --
-    Scan.optional (Args.$$$ "for" |-- Scan.repeat1 Args.name) []) >>
-  (fn (raw_props, fixes) => fn ctxt =>
-    let
-      val (_, ctxt') = Variable.add_fixes fixes ctxt;
-      val props = Syntax.read_props ctxt' raw_props;
-      val ctxt'' = fold Variable.declare_term props ctxt';
-      val rules = ProofContext.export ctxt'' ctxt (map (mk_cases ctxt'') props)
-    in Method.erule 0 rules end);
-
+val ind_cases_setup =
+  Method.setup @{binding ind_cases}
+    (Scan.lift (Scan.repeat1 Args.name_source --
+      Scan.optional (Args.$$$ "for" |-- Scan.repeat1 Args.name) []) >>
+      (fn (raw_props, fixes) => fn ctxt =>
+        let
+          val (_, ctxt') = Variable.add_fixes fixes ctxt;
+          val props = Syntax.read_props ctxt' raw_props;
+          val ctxt'' = fold Variable.declare_term props ctxt';
+          val rules = ProofContext.export ctxt'' ctxt (map (mk_cases ctxt'') props)
+        in Method.erule 0 rules end))
+    "dynamic case analysis on predicates";
 
 
 (* prove induction rule *)
@@ -934,7 +935,7 @@
 (* setup theory *)
 
 val setup =
-  Method.setup @{binding ind_cases} ind_cases "dynamic case analysis on predicates" #>
+  ind_cases_setup #>
   Attrib.setup @{binding mono} (Attrib.add_del mono_add mono_del)
     "declaration of monotonicity rule";
 
--- a/src/HOL/Tools/inductive_realizer.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/inductive_realizer.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -505,13 +505,11 @@
       | SOME (SOME sets') => sets \\ sets')
   end I);
 
-val ind_realizer =
-  (Scan.option (Scan.lift (Args.$$$ "irrelevant") |--
-    Scan.option (Scan.lift (Args.colon) |--
-      Scan.repeat1 Args.const))) >> rlz_attrib;
-
 val setup =
-  Attrib.setup @{binding ind_realizer} ind_realizer "add realizers for inductive set";
+  Attrib.setup @{binding ind_realizer}
+    ((Scan.option (Scan.lift (Args.$$$ "irrelevant") |--
+      Scan.option (Scan.lift (Args.colon) |-- Scan.repeat1 Args.const))) >> rlz_attrib)
+    "add realizers for inductive set";
 
 end;
 
--- a/src/HOL/Tools/int_arith.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/int_arith.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -377,6 +377,8 @@
 struct
   val assoc_ss = HOL_ss addsimps @{thms mult_ac}
   val eq_reflection = eq_reflection
+  fun is_numeral (Const(@{const_name Int.number_of}, _) $ _) = true
+    | is_numeral _ = false;
 end;
 
 structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data);
@@ -530,7 +532,7 @@
   :: Int_Numeral_Simprocs.cancel_numerals;
 
 val setup =
-  LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
+  Lin_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
    {add_mono_thms = add_mono_thms,
     mult_mono_thms = @{thm mult_strict_left_mono} :: @{thm mult_left_mono} :: mult_mono_thms,
     inj_thms = nat_inj_thms @ inj_thms,
@@ -547,7 +549,7 @@
   "fast_int_arith" 
      ["(m::'a::{ordered_idom,number_ring}) < n",
       "(m::'a::{ordered_idom,number_ring}) <= n",
-      "(m::'a::{ordered_idom,number_ring}) = n"] (K LinArith.lin_arith_simproc);
+      "(m::'a::{ordered_idom,number_ring}) = n"] (K Lin_Arith.lin_arith_simproc);
 
 end;
 
--- a/src/HOL/Tools/int_factor_simprocs.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/int_factor_simprocs.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -218,9 +218,30 @@
 val simplify_one = Arith_Data.simplify_meta_eq  
   [@{thm mult_1_left}, @{thm mult_1_right}, @{thm div_by_1}, @{thm numeral_1_eq_1}];
 
-fun cancel_simplify_meta_eq cancel_th ss th =
+fun cancel_simplify_meta_eq ss cancel_th th =
     simplify_one ss (([th, cancel_th]) MRS trans);
 
+local
+  val Tp_Eq = Thm.reflexive(Thm.cterm_of (@{theory HOL}) HOLogic.Trueprop)
+  fun Eq_True_elim Eq = 
+    Thm.equal_elim (Thm.combination Tp_Eq (Thm.symmetric Eq)) @{thm TrueI}
+in
+fun sign_conv pos_th neg_th ss t =
+  let val T = fastype_of t;
+      val zero = Const(@{const_name HOL.zero}, T);
+      val less = Const(@{const_name HOL.less}, [T,T] ---> HOLogic.boolT);
+      val pos = less $ zero $ t and neg = less $ t $ zero
+      fun prove p =
+        Option.map Eq_True_elim (Lin_Arith.lin_arith_simproc ss p)
+        handle THM _ => NONE
+    in case prove pos of
+         SOME th => SOME(th RS pos_th)
+       | NONE => (case prove neg of
+                    SOME th => SOME(th RS neg_th)
+                  | NONE => NONE)
+    end;
+end
+
 structure CancelFactorCommon =
   struct
   val mk_sum            = long_mk_prod
@@ -231,6 +252,7 @@
   val trans_tac         = K Arith_Data.trans_tac
   val norm_ss = HOL_ss addsimps mult_1s @ @{thms mult_ac}
   fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss))
+  val simplify_meta_eq  = cancel_simplify_meta_eq 
   end;
 
 (*mult_cancel_left requires a ring with no zero divisors.*)
@@ -239,7 +261,27 @@
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_eq
   val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
-  val simplify_meta_eq  = cancel_simplify_meta_eq @{thm mult_cancel_left}
+  val simp_conv = K (K (SOME @{thm mult_cancel_left}))
+);
+
+(*for ordered rings*)
+structure LeCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} Term.dummyT
+  val simp_conv = sign_conv
+    @{thm mult_le_cancel_left_pos} @{thm mult_le_cancel_left_neg}
+);
+
+(*for ordered rings*)
+structure LessCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.less} Term.dummyT
+  val simp_conv = sign_conv
+    @{thm mult_less_cancel_left_pos} @{thm mult_less_cancel_left_neg}
 );
 
 (*zdiv_zmult_zmult1_if is for integer division (div).*)
@@ -248,7 +290,7 @@
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
   val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.intT
-  val simplify_meta_eq  = cancel_simplify_meta_eq @{thm zdiv_zmult_zmult1_if}
+  val simp_conv = K (K (SOME @{thm zdiv_zmult_zmult1_if}))
 );
 
 structure IntModCancelFactor = ExtractCommonTermFun
@@ -256,7 +298,7 @@
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_binop @{const_name Divides.mod}
   val dest_bal = HOLogic.dest_bin @{const_name Divides.mod} HOLogic.intT
-  val simplify_meta_eq  = cancel_simplify_meta_eq @{thm zmod_zmult_zmult1}
+  val simp_conv = K (K (SOME @{thm zmod_zmult_zmult1}))
 );
 
 structure IntDvdCancelFactor = ExtractCommonTermFun
@@ -264,7 +306,7 @@
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd}
   val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} Term.dummyT
-  val simplify_meta_eq  = cancel_simplify_meta_eq @{thm dvd_mult_cancel_left}
+  val simp_conv = K (K (SOME @{thm dvd_mult_cancel_left}))
 );
 
 (*Version for all fields, including unordered ones (type complex).*)
@@ -273,7 +315,7 @@
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_binop @{const_name HOL.divide}
   val dest_bal = HOLogic.dest_bin @{const_name HOL.divide} Term.dummyT
-  val simplify_meta_eq  = cancel_simplify_meta_eq @{thm mult_divide_mult_cancel_left_if}
+  val simp_conv = K (K (SOME @{thm mult_divide_mult_cancel_left_if}))
 );
 
 val cancel_factors =
@@ -281,7 +323,15 @@
    [("ring_eq_cancel_factor",
      ["(l::'a::{idom}) * m = n",
       "(l::'a::{idom}) = m * n"],
-    K EqCancelFactor.proc),
+     K EqCancelFactor.proc),
+    ("ordered_ring_le_cancel_factor",
+     ["(l::'a::ordered_ring) * m <= n",
+      "(l::'a::ordered_ring) <= m * n"],
+     K LeCancelFactor.proc),
+    ("ordered_ring_less_cancel_factor",
+     ["(l::'a::ordered_ring) * m < n",
+      "(l::'a::ordered_ring) < m * n"],
+     K LessCancelFactor.proc),
     ("int_div_cancel_factor",
      ["((l::int) * m) div n", "(l::int) div (m * n)"],
      K IntDivCancelFactor.proc),
--- a/src/HOL/Tools/lin_arith.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/lin_arith.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -6,13 +6,9 @@
 
 signature BASIC_LIN_ARITH =
 sig
-  type arith_tactic
-  val mk_arith_tactic: string -> (Proof.context -> int -> tactic) -> arith_tactic
-  val eq_arith_tactic: arith_tactic * arith_tactic -> bool
   val arith_split_add: attribute
   val arith_discrete: string -> Context.generic -> Context.generic
   val arith_inj_const: string * typ -> Context.generic -> Context.generic
-  val arith_tactic_add: arith_tactic -> Context.generic -> Context.generic
   val fast_arith_split_limit: int Config.T
   val fast_arith_neq_limit: int Config.T
   val lin_arith_pre_tac: Proof.context -> int -> tactic
@@ -21,9 +17,7 @@
   val trace_arith: bool ref
   val lin_arith_simproc: simpset -> term -> thm option
   val fast_nat_arith_simproc: simproc
-  val simple_arith_tac: Proof.context -> int -> tactic
-  val arith_tac: Proof.context -> int -> tactic
-  val silent_arith_tac: Proof.context -> int -> tactic
+  val linear_arith_tac: Proof.context -> int -> tactic
 end;
 
 signature LIN_ARITH =
@@ -39,7 +33,7 @@
   val setup: Context.generic -> Context.generic
 end;
 
-structure LinArith: LIN_ARITH =
+structure Lin_Arith: LIN_ARITH =
 struct
 
 (* Parameters data for general linear arithmetic functor *)
@@ -72,7 +66,7 @@
   let val _ $ t = Thm.prop_of thm
   in t = Const("False",HOLogic.boolT) end;
 
-fun is_nat(t) = fastype_of1 t = HOLogic.natT;
+fun is_nat t = (fastype_of1 t = HOLogic.natT);
 
 fun mk_nat_thm sg t =
   let val ct = cterm_of sg t  and cn = cterm_of sg (Var(("n",0),HOLogic.natT))
@@ -83,49 +77,35 @@
 
 (* arith context data *)
 
-datatype arith_tactic =
-  ArithTactic of {name: string, tactic: Proof.context -> int -> tactic, id: stamp};
-
-fun mk_arith_tactic name tactic = ArithTactic {name = name, tactic = tactic, id = stamp ()};
-
-fun eq_arith_tactic (ArithTactic {id = id1, ...}, ArithTactic {id = id2, ...}) = (id1 = id2);
-
 structure ArithContextData = GenericDataFun
 (
   type T = {splits: thm list,
             inj_consts: (string * typ) list,
-            discrete: string list,
-            tactics: arith_tactic list};
-  val empty = {splits = [], inj_consts = [], discrete = [], tactics = []};
+            discrete: string list};
+  val empty = {splits = [], inj_consts = [], discrete = []};
   val extend = I;
   fun merge _
-   ({splits= splits1, inj_consts= inj_consts1, discrete= discrete1, tactics= tactics1},
-    {splits= splits2, inj_consts= inj_consts2, discrete= discrete2, tactics= tactics2}) : T =
+   ({splits= splits1, inj_consts= inj_consts1, discrete= discrete1},
+    {splits= splits2, inj_consts= inj_consts2, discrete= discrete2}) : T =
    {splits = Library.merge Thm.eq_thm_prop (splits1, splits2),
     inj_consts = Library.merge (op =) (inj_consts1, inj_consts2),
-    discrete = Library.merge (op =) (discrete1, discrete2),
-    tactics = Library.merge eq_arith_tactic (tactics1, tactics2)};
+    discrete = Library.merge (op =) (discrete1, discrete2)};
 );
 
 val get_arith_data = ArithContextData.get o Context.Proof;
 
 val arith_split_add = Thm.declaration_attribute (fn thm =>
-  ArithContextData.map (fn {splits, inj_consts, discrete, tactics} =>
+  ArithContextData.map (fn {splits, inj_consts, discrete} =>
     {splits = update Thm.eq_thm_prop thm splits,
-     inj_consts = inj_consts, discrete = discrete, tactics = tactics}));
-
-fun arith_discrete d = ArithContextData.map (fn {splits, inj_consts, discrete, tactics} =>
-  {splits = splits, inj_consts = inj_consts,
-   discrete = update (op =) d discrete, tactics = tactics});
+     inj_consts = inj_consts, discrete = discrete}));
 
-fun arith_inj_const c = ArithContextData.map (fn {splits, inj_consts, discrete, tactics} =>
-  {splits = splits, inj_consts = update (op =) c inj_consts,
-   discrete = discrete, tactics= tactics});
+fun arith_discrete d = ArithContextData.map (fn {splits, inj_consts, discrete} =>
+  {splits = splits, inj_consts = inj_consts,
+   discrete = update (op =) d discrete});
 
-fun arith_tactic_add tac = ArithContextData.map (fn {splits, inj_consts, discrete, tactics} =>
-  {splits = splits, inj_consts = inj_consts, discrete = discrete,
-   tactics = update eq_arith_tactic tac tactics});
-
+fun arith_inj_const c = ArithContextData.map (fn {splits, inj_consts, discrete} =>
+  {splits = splits, inj_consts = update (op =) c inj_consts,
+   discrete = discrete});
 
 val (fast_arith_split_limit, setup1) = Attrib.config_int "fast_arith_split_limit" 9;
 val (fast_arith_neq_limit, setup2) = Attrib.config_int "fast_arith_neq_limit" 9;
@@ -794,7 +774,7 @@
    Most of the work is done by the cancel tactics. *)
 
 val init_arith_data =
- Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, ...} =>
+ map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, ...} =>
    {add_mono_thms = add_mono_thms @
     @{thms add_mono_thms_ordered_semiring} @ @{thms add_mono_thms_ordered_field},
     mult_mono_thms = mult_mono_thms,
@@ -815,7 +795,7 @@
   arith_discrete "nat";
 
 fun add_arith_facts ss =
-  add_prems (ArithFacts.get (MetaSimplifier.the_context ss)) ss;
+  add_prems (Arith_Data.get_arith_facts (MetaSimplifier.the_context ss)) ss;
 
 val lin_arith_simproc = add_arith_facts #> Fast_Arith.lin_arith_simproc;
 
@@ -895,27 +875,12 @@
     (REPEAT_DETERM o split_tac (#splits (get_arith_data ctxt)))
     (fast_ex_arith_tac ctxt ex);
 
-fun more_arith_tacs ctxt =
-  let val tactics = #tactics (get_arith_data ctxt)
-  in FIRST' (map (fn ArithTactic {tactic, ...} => tactic ctxt) tactics) end;
-
 in
 
-fun simple_arith_tac ctxt = FIRST' [fast_arith_tac ctxt,
-  ObjectLogic.full_atomize_tac THEN' (REPEAT_DETERM o rtac impI) THEN' raw_arith_tac ctxt true];
-
-fun arith_tac ctxt = FIRST' [fast_arith_tac ctxt,
-  ObjectLogic.full_atomize_tac THEN' (REPEAT_DETERM o rtac impI) THEN' raw_arith_tac ctxt true,
-  more_arith_tacs ctxt];
+fun gen_linear_arith_tac ex ctxt = FIRST' [fast_arith_tac ctxt,
+  ObjectLogic.full_atomize_tac THEN' (REPEAT_DETERM o rtac impI) THEN' raw_arith_tac ctxt ex];
 
-fun silent_arith_tac ctxt = FIRST' [fast_arith_tac ctxt,
-  ObjectLogic.full_atomize_tac THEN' (REPEAT_DETERM o rtac impI) THEN' raw_arith_tac ctxt false,
-  more_arith_tacs ctxt];
-
-val arith_method = Args.bang_facts >>
-  (fn prems => fn ctxt => METHOD (fn facts =>
-      HEADGOAL (Method.insert_tac (prems @ ArithFacts.get ctxt @ facts)
-      THEN' arith_tac ctxt)));
+val linear_arith_tac = gen_linear_arith_tac true;
 
 end;
 
@@ -929,11 +894,16 @@
       (add_arith_facts #> Fast_Arith.cut_lin_arith_tac))) #>
   Context.mapping
    (setup_options #>
-    Method.setup @{binding arith} arith_method "decide linear arithmetic" #>
+    Arith_Data.add_tactic "linear arithmetic" gen_linear_arith_tac #>
+    Method.setup @{binding linarith}
+      (Args.bang_facts >> (fn prems => fn ctxt =>
+        METHOD (fn facts =>
+          HEADGOAL (Method.insert_tac (prems @ Arith_Data.get_arith_facts ctxt @ facts)
+            THEN' linear_arith_tac ctxt)))) "linear arithmetic" #>
     Attrib.setup @{binding arith_split} (Scan.succeed arith_split_add)
       "declaration of split rules for arithmetic procedure") I;
 
 end;
 
-structure BasicLinArith: BASIC_LIN_ARITH = LinArith;
-open BasicLinArith;
+structure Basic_Lin_Arith: BASIC_LIN_ARITH = Lin_Arith;
+open Basic_Lin_Arith;
--- a/src/HOL/Tools/meson.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/meson.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -25,7 +25,7 @@
   val skolemize_prems_tac: thm list -> int -> tactic
   val MESON: (thm list -> thm list) -> (thm list -> tactic) -> int -> tactic
   val best_meson_tac: (thm -> int) -> int -> tactic
-  val safe_best_meson_tac: int -> tactic
+  val safe_best_meson_tac: claset -> int -> tactic
   val depth_meson_tac: int -> tactic
   val prolog_step_tac': thm list -> int -> tactic
   val iter_deepen_prolog_tac: thm list -> tactic
@@ -33,7 +33,7 @@
   val make_meta_clause: thm -> thm
   val make_meta_clauses: thm list -> thm list
   val meson_claset_tac: thm list -> claset -> int -> tactic
-  val meson_tac: int -> tactic
+  val meson_tac: claset -> int -> tactic
   val negate_head: thm -> thm
   val select_literal: int -> thm -> thm
   val skolemize_tac: int -> tactic
@@ -589,8 +589,8 @@
                          (prolog_step_tac (make_horns cls) 1));
 
 (*First, breaks the goal into independent units*)
-val safe_best_meson_tac =
-     SELECT_GOAL (TRY (CLASET safe_tac) THEN
+fun safe_best_meson_tac cs =
+     SELECT_GOAL (TRY (safe_tac cs) THEN
                   TRYALL (best_meson_tac size_of_subgoals));
 
 (** Depth-first search version **)
@@ -634,7 +634,7 @@
 fun meson_claset_tac ths cs =
   SELECT_GOAL (TRY (safe_tac cs) THEN TRYALL (iter_deepen_meson_tac ths));
 
-val meson_tac = CLASET' (meson_claset_tac []);
+val meson_tac = meson_claset_tac [];
 
 
 (**** Code to support ordinary resolution, rather than Model Elimination ****)
--- a/src/HOL/Tools/nat_simprocs.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/nat_simprocs.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -352,7 +352,7 @@
 val simplify_one = Arith_Data.simplify_meta_eq  
   [@{thm mult_1_left}, @{thm mult_1_right}, @{thm div_1}, @{thm numeral_1_eq_Suc_0}];
 
-fun cancel_simplify_meta_eq cancel_th ss th =
+fun cancel_simplify_meta_eq ss cancel_th th =
     simplify_one ss (([th, cancel_th]) MRS trans);
 
 structure CancelFactorCommon =
@@ -365,6 +365,7 @@
   val trans_tac         = K Arith_Data.trans_tac
   val norm_ss = HOL_ss addsimps mult_1s @ @{thms mult_ac}
   fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss))
+  val simplify_meta_eq  = cancel_simplify_meta_eq
   end;
 
 structure EqCancelFactor = ExtractCommonTermFun
@@ -372,7 +373,7 @@
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_eq
   val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
-  val simplify_meta_eq  = cancel_simplify_meta_eq @{thm nat_mult_eq_cancel_disj}
+  val simp_conv = K(K (SOME @{thm nat_mult_eq_cancel_disj}))
 );
 
 structure LessCancelFactor = ExtractCommonTermFun
@@ -380,7 +381,7 @@
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
   val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT
-  val simplify_meta_eq  = cancel_simplify_meta_eq @{thm nat_mult_less_cancel_disj}
+  val simp_conv = K(K (SOME @{thm nat_mult_less_cancel_disj}))
 );
 
 structure LeCancelFactor = ExtractCommonTermFun
@@ -388,7 +389,7 @@
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
   val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT
-  val simplify_meta_eq  = cancel_simplify_meta_eq @{thm nat_mult_le_cancel_disj}
+  val simp_conv = K(K (SOME @{thm nat_mult_le_cancel_disj}))
 );
 
 structure DivideCancelFactor = ExtractCommonTermFun
@@ -396,7 +397,7 @@
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
   val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.natT
-  val simplify_meta_eq  = cancel_simplify_meta_eq @{thm nat_mult_div_cancel_disj}
+  val simp_conv = K(K (SOME @{thm nat_mult_div_cancel_disj}))
 );
 
 structure DvdCancelFactor = ExtractCommonTermFun
@@ -404,7 +405,7 @@
   val prove_conv = Arith_Data.prove_conv
   val mk_bal   = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd}
   val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.natT
-  val simplify_meta_eq  = cancel_simplify_meta_eq @{thm nat_mult_dvd_cancel_disj}
+  val simp_conv = K(K (SOME @{thm nat_mult_dvd_cancel_disj}))
 );
 
 val cancel_factor =
@@ -564,7 +565,7 @@
 in
 
 val nat_simprocs_setup =
-  LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
+  Lin_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
    {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms,
     inj_thms = inj_thms, lessD = lessD, neqE = neqE,
     simpset = simpset addsimps add_rules
--- a/src/HOL/Tools/rat_arith.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/rat_arith.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -35,7 +35,7 @@
 in
 
 val rat_arith_setup =
-  LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
+  Lin_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
    {add_mono_thms = add_mono_thms,
     mult_mono_thms = mult_mono_thms,
     inj_thms = int_inj_thms @ nat_inj_thms @ inj_thms,
--- a/src/HOL/Tools/real_arith.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/real_arith.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -29,7 +29,7 @@
 in
 
 val real_arith_setup =
-  LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
+  Lin_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
    {add_mono_thms = add_mono_thms,
     mult_mono_thms = mult_mono_thms,
     inj_thms = int_inj_thms @ nat_inj_thms @ inj_thms,
--- a/src/HOL/Tools/record_package.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/record_package.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -1782,7 +1782,8 @@
     val alphas_fields = List.foldr OldTerm.add_typ_tfree_names [] types;
     val alphas_ext = alphas inter alphas_fields;
     val len = length fields;
-    val variants = Name.variant_list (moreN::rN::rN ^ "'"::wN::parent_variants) (map fst bfields);
+    val variants =
+      Name.variant_list (moreN :: rN :: (rN ^ "'") :: wN :: parent_variants) (map fst bfields);
     val vars = ListPair.map Free (variants, types);
     val named_vars = names ~~ vars;
     val idxs = 0 upto (len - 1);
--- a/src/HOL/Tools/res_axioms.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/res_axioms.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -521,18 +521,19 @@
           REPEAT_DETERM_N (length ts) o (etac thin_rl)]
      end);
 
-val setup_methods =
+val neg_clausify_setup =
   Method.setup @{binding neg_clausify} (Scan.succeed (K (SIMPLE_METHOD' neg_clausify_tac)))
   "conversion of goal to conjecture clauses";
 
 
 (** Attribute for converting a theorem into clauses **)
 
-val clausify = Scan.lift OuterParse.nat >>
-  (fn i => Thm.rule_attribute (fn context => fn th =>
-      Meson.make_meta_clause (nth (cnf_axiom (Context.theory_of context) th) i)));
-
-val setup_attrs = Attrib.setup @{binding clausify} clausify "conversion of theorem to clauses";
+val clausify_setup =
+  Attrib.setup @{binding clausify}
+    (Scan.lift OuterParse.nat >>
+      (fn i => Thm.rule_attribute (fn context => fn th =>
+          Meson.make_meta_clause (nth (cnf_axiom (Context.theory_of context) th) i))))
+  "conversion of theorem to clauses";
 
 
 
@@ -540,8 +541,8 @@
 
 val setup =
   meson_method_setup #>
-  setup_methods #>
-  setup_attrs #>
+  neg_clausify_setup #>
+  clausify_setup #>
   perhaps saturate_skolem_cache #>
   Theory.at_end clause_cache_endtheory;
 
--- a/src/HOL/Tools/simpdata.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/simpdata.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -147,8 +147,6 @@
 
 val op addsplits     = Splitter.addsplits;
 val op delsplits     = Splitter.delsplits;
-val Addsplits        = Splitter.Addsplits;
-val Delsplits        = Splitter.Delsplits;
 
 
 (* integration of simplifier with classical reasoner *)
--- a/src/HOL/Tools/split_rule.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Tools/split_rule.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -141,18 +141,18 @@
     |> RuleCases.save rl
   end;
 
+
 (* attribute syntax *)
 
 (* FIXME dynamically scoped due to Args.name_source/pair_tac *)
 
-val split_format = Scan.lift
- (Args.parens (Args.$$$ "complete") >> K (Thm.rule_attribute (K complete_split_rule)) ||
-  OuterParse.and_list1 (Scan.repeat Args.name_source)
-    >> (fn xss => Thm.rule_attribute (fn context =>
-        split_rule_goal (Context.proof_of context) xss)));
-
 val setup =
-  Attrib.setup @{binding split_format} split_format
+  Attrib.setup @{binding split_format}
+    (Scan.lift
+     (Args.parens (Args.$$$ "complete") >> K (Thm.rule_attribute (K complete_split_rule)) ||
+      OuterParse.and_list1 (Scan.repeat Args.name_source)
+        >> (fn xss => Thm.rule_attribute (fn context =>
+            split_rule_goal (Context.proof_of context) xss))))
     "split pair-typed subterms in premises, or function arguments" #>
   Attrib.setup @{binding split_rule} (Scan.succeed (Thm.rule_attribute (K split_rule)))
     "curries ALL function variables occurring in a rule's conclusion";
--- a/src/HOL/UNITY/UNITY_tactics.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/UNITY/UNITY_tactics.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -59,6 +59,6 @@
      th RS @{thm o_equiv_assoc} |> simplify (HOL_ss addsimps [@{thm o_assoc}]),
      th RS @{thm o_equiv_apply} |> simplify (HOL_ss addsimps [@{thm o_def}, @{thm sub_def}])];
 
-Addsimps (make_o_equivs @{thm fst_o_funPair} @ make_o_equivs @{thm snd_o_funPair});
-
-Addsimps (make_o_equivs @{thm fst_o_lift_map} @ make_o_equivs @{thm snd_o_lift_map});
+Simplifier.change_simpset (fn ss => ss
+  addsimps (make_o_equivs @{thm fst_o_funPair} @ make_o_equivs @{thm snd_o_funPair})
+  addsimps (make_o_equivs @{thm fst_o_lift_map} @ make_o_equivs @{thm snd_o_lift_map}));
--- a/src/HOL/Word/TdThs.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Word/TdThs.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -89,7 +89,7 @@
 
 end
 
-interpretation nat_int!: type_definition int nat "Collect (op <= 0)"
+interpretation nat_int: type_definition int nat "Collect (op <= 0)"
   by (rule td_nat_int)
 
 declare
--- a/src/HOL/Word/WordArith.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Word/WordArith.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -21,7 +21,7 @@
 proof
 qed (unfold word_sle_def word_sless_def, auto)
 
-interpretation signed!: linorder "word_sle" "word_sless"
+interpretation signed: linorder "word_sle" "word_sless"
   by (rule signed_linorder)
 
 lemmas word_arith_wis = 
@@ -511,10 +511,13 @@
      addcongs @{thms power_False_cong}
 
 fun uint_arith_tacs ctxt = 
-  let fun arith_tac' n t = arith_tac ctxt n t handle COOPER => Seq.empty  
+  let
+    fun arith_tac' n t = Arith_Data.verbose_arith_tac ctxt n t handle COOPER => Seq.empty;
+    val cs = local_claset_of ctxt;
+    val ss = local_simpset_of ctxt;
   in 
-    [ CLASET' clarify_tac 1,
-      SIMPSET' (full_simp_tac o uint_arith_ss_of) 1,
+    [ clarify_tac cs 1,
+      full_simp_tac (uint_arith_ss_of ss) 1,
       ALLGOALS (full_simp_tac (HOL_ss addsplits @{thms uint_splits} 
                                       addcongs @{thms power_False_cong})),
       rewrite_goals_tac @{thms word_size}, 
@@ -698,6 +701,7 @@
   apply (erule (2) udvd_decr0)
   done
 
+ML{*Delsimprocs cancel_factors*}
 lemma udvd_incr2_K: 
   "p < a + s ==> a <= a + s ==> K udvd s ==> K udvd p - a ==> a <= p ==> 
     0 < K ==> p <= p + K & p + K <= a + s"
@@ -713,6 +717,7 @@
    apply arith
   apply simp
   done
+ML{*Delsimprocs cancel_factors*}
 
 (* links with rbl operations *)
 lemma word_succ_rbl:
@@ -857,7 +862,7 @@
 lemmas td_ext_unat = refl [THEN td_ext_unat']
 lemmas unat_of_nat = td_ext_unat [THEN td_ext.eq_norm, standard]
 
-interpretation word_unat!:
+interpretation word_unat:
   td_ext "unat::'a::len word => nat" 
          of_nat 
          "unats (len_of TYPE('a::len))"
@@ -1069,10 +1074,13 @@
      addcongs @{thms power_False_cong}
 
 fun unat_arith_tacs ctxt =   
-  let fun arith_tac' n t = arith_tac ctxt n t handle COOPER => Seq.empty  
+  let
+    fun arith_tac' n t = Arith_Data.verbose_arith_tac ctxt n t handle COOPER => Seq.empty;
+    val cs = local_claset_of ctxt;
+    val ss = local_simpset_of ctxt;
   in 
-    [ CLASET' clarify_tac 1,
-      SIMPSET' (full_simp_tac o unat_arith_ss_of) 1,
+    [ clarify_tac cs 1,
+      full_simp_tac (unat_arith_ss_of ss) 1,
       ALLGOALS (full_simp_tac (HOL_ss addsplits @{thms unat_splits} 
                                        addcongs @{thms power_False_cong})),
       rewrite_goals_tac @{thms word_size}, 
--- a/src/HOL/Word/WordBitwise.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Word/WordBitwise.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -343,7 +343,7 @@
 
 lemmas td_ext_nth = td_ext_nth' [OF refl refl refl, unfolded word_size]
 
-interpretation test_bit!:
+interpretation test_bit:
   td_ext "op !! :: 'a::len0 word => nat => bool"
          set_bits
          "{f. \<forall>i. f i \<longrightarrow> i < len_of TYPE('a::len0)}"
--- a/src/HOL/Word/WordDefinition.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Word/WordDefinition.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -351,7 +351,7 @@
 
 lemmas int_word_uint = td_ext_uint [THEN td_ext.eq_norm, standard]
 
-interpretation word_uint!:
+interpretation word_uint:
   td_ext "uint::'a::len0 word \<Rightarrow> int" 
          word_of_int 
          "uints (len_of TYPE('a::len0))"
@@ -363,7 +363,7 @@
 lemmas td_ext_ubin = td_ext_uint 
   [simplified len_gt_0 no_bintr_alt1 [symmetric]]
 
-interpretation word_ubin!:
+interpretation word_ubin:
   td_ext "uint::'a::len0 word \<Rightarrow> int" 
          word_of_int 
          "uints (len_of TYPE('a::len0))"
@@ -418,7 +418,7 @@
    and interpretations do not produce thm duplicates. I.e. 
    we get the name word_sint.Rep_eqD, but not word_sbin.Req_eqD,
    because the latter is the same thm as the former *)
-interpretation word_sint!:
+interpretation word_sint:
   td_ext "sint ::'a::len word => int" 
           word_of_int 
           "sints (len_of TYPE('a::len))"
@@ -426,7 +426,7 @@
                2 ^ (len_of TYPE('a::len) - 1)"
   by (rule td_ext_sint)
 
-interpretation word_sbin!:
+interpretation word_sbin:
   td_ext "sint ::'a::len word => int" 
           word_of_int 
           "sints (len_of TYPE('a::len))"
@@ -630,7 +630,7 @@
   apply simp
   done
 
-interpretation word_bl!:
+interpretation word_bl:
   type_definition "to_bl :: 'a::len0 word => bool list"
                   of_bl  
                   "{bl. length bl = len_of TYPE('a::len0)}"
--- a/src/HOL/Word/WordGenLib.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/Word/WordGenLib.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -106,7 +106,7 @@
   apply (rule word_or_not)
   done
 
-interpretation word_bool_alg!:
+interpretation word_bool_alg:
   boolean "op AND" "op OR" bitNOT 0 max_word
   by (rule word_boolean)
 
@@ -114,7 +114,7 @@
   "x XOR y = x AND NOT y OR NOT x AND (y::'a::len0 word)"
   by (rule word_eqI) (auto simp add: word_ops_nth_size word_size)
 
-interpretation word_bool_alg!:
+interpretation word_bool_alg:
   boolean_xor "op AND" "op OR" bitNOT 0 max_word "op XOR"
   apply (rule boolean_xor.intro)
    apply (rule word_boolean)
--- a/src/HOL/ex/Arith_Examples.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/ex/Arith_Examples.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:  HOL/ex/Arith_Examples.thy
-    ID:     $Id$
     Author: Tjark Weber
 *)
 
@@ -14,13 +13,13 @@
 
   @{ML fast_arith_tac} is a very basic version of the tactic.  It performs no
   meta-to-object-logic conversion, and only some splitting of operators.
-  @{ML simple_arith_tac} performs meta-to-object-logic conversion, full
+  @{ML linear_arith_tac} performs meta-to-object-logic conversion, full
   splitting of operators, and NNF normalization of the goal.  The @{text arith}
   method combines them both, and tries other methods (e.g.~@{text presburger})
   as well.  This is the one that you should use in your proofs!
 
   An @{text arith}-based simproc is available as well (see @{ML
-  LinArith.lin_arith_simproc}), which---for performance
+  Lin_Arith.lin_arith_simproc}), which---for performance
   reasons---however does even less splitting than @{ML fast_arith_tac}
   at the moment (namely inequalities only).  (On the other hand, it
   does take apart conjunctions, which @{ML fast_arith_tac} currently
@@ -83,7 +82,7 @@
   by (tactic {* fast_arith_tac @{context} 1 *})
 
 lemma "!!x. ((x::nat) <= y) = (x - y = 0)"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 lemma "[| (x::nat) < y; d < 1 |] ==> x - y = d"
   by (tactic {* fast_arith_tac @{context} 1 *})
@@ -140,34 +139,34 @@
 subsection {* Meta-Logic *}
 
 lemma "x < Suc y == x <= y"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 lemma "((x::nat) == z ==> x ~= y) ==> x ~= y | z ~= y"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 
 subsection {* Various Other Examples *}
 
 lemma "(x < Suc y) = (x <= y)"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 lemma "[| (x::nat) < y; y < z |] ==> x < z"
   by (tactic {* fast_arith_tac @{context} 1 *})
 
 lemma "(x::nat) < y & y < z ==> x < z"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 text {* This example involves no arithmetic at all, but is solved by
   preprocessing (i.e. NNF normalization) alone. *}
 
 lemma "(P::bool) = Q ==> Q = P"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 lemma "[| P = (x = 0); (~P) = (y = 0) |] ==> min (x::nat) y = 0"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 lemma "[| P = (x = 0); (~P) = (y = 0) |] ==> max (x::nat) y = x + y"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 lemma "[| (x::nat) ~= y; a + 2 = b; a < y; y < b; a < x; x < b |] ==> False"
   by (tactic {* fast_arith_tac @{context} 1 *})
@@ -185,7 +184,7 @@
   by (tactic {* fast_arith_tac @{context} 1 *})
 
 lemma "[| (x::nat) < y; P (x - y) |] ==> P 0"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 lemma "(x - y) - (x::nat) = (x - x) - y"
   by (tactic {* fast_arith_tac @{context} 1 *})
@@ -207,7 +206,7 @@
 (*        preprocessing negates the goal and tries to compute its negation *)
 (*        normal form, which creates lots of separate cases for this       *)
 (*        disjunction of conjunctions                                      *)
-(* by (tactic {* simple_arith_tac 1 *}) *)
+(* by (tactic {* linear_arith_tac 1 *}) *)
 oops
 
 lemma "2 * (x::nat) ~= 1"
--- a/src/HOL/ex/Classical.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/ex/Classical.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/ex/Classical
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1994  University of Cambridge
 *)
@@ -430,7 +429,7 @@
 lemma "(\<forall>x y z. R(x,y) & R(y,z) --> R(x,z)) &
        (\<forall>x. \<exists>y. R(x,y)) -->
        ~ (\<forall>x. P x = (\<forall>y. R(x,y) --> ~ P y))"
-by (tactic{*Meson.safe_best_meson_tac 1*})
+by (tactic{*Meson.safe_best_meson_tac @{claset} 1*})
     --{*In contrast, @{text meson} is SLOW: 7.6s on griffon*}
 
 
@@ -724,7 +723,7 @@
        (\<forall>x y. bird x & snail y \<longrightarrow> ~eats x y) &
        (\<forall>x. (caterpillar x \<or> snail x) \<longrightarrow> (\<exists>y. plant y & eats x y))
        \<longrightarrow> (\<exists>x y. animal x & animal y & (\<exists>z. grain z & eats y z & eats x y))"
-by (tactic{*Meson.safe_best_meson_tac 1*})
+by (tactic{*Meson.safe_best_meson_tac @{claset} 1*})
     --{*Nearly twice as fast as @{text meson},
         which performs iterative deepening rather than best-first search*}
 
--- a/src/HOL/ex/Code_Antiq.thy	Sat Mar 28 00:11:02 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,84 +0,0 @@
-(*  Title:      HOL/ex/Code_Antiq.thy
-    ID:         $Id$
-    Author:     Florian Haftmann
-*)
-
-header {* A simple certificate check as toy example for the code antiquotation *}
-
-theory Code_Antiq
-imports Plain
-begin
-
-lemma div_cert1:
-  fixes n m q r :: nat
-  assumes "r < m"
-    and "0 < m"
-    and "n = m * q + r"
-  shows "n div m = q"
-using assms by (simp add: div_mult_self2 add_commute)
-
-lemma div_cert2:
-  fixes n :: nat
-  shows "n div 0 = 0"
-by simp
-
-ML {*
-local
-
-fun code_of_val k = if k <= 0 then @{code "0::nat"}
-  else @{code Suc} (code_of_val (k - 1));
-
-fun val_of_code @{code "0::nat"} = 0
-  | val_of_code (@{code Suc} n) = val_of_code n + 1;
-
-val term_of_code = HOLogic.mk_nat o val_of_code;
-
-infix 9 &$;
-val op &$ = uncurry Thm.capply;
-
-val simpset = HOL_ss addsimps
-  @{thms plus_nat.simps add_0_right add_Suc_right times_nat.simps mult_0_right mult_Suc_right  less_nat_zero_code le_simps(2) less_eq_nat.simps(1) le_simps(3)}
-
-fun prove prop = Goal.prove_internal [] (@{cterm Trueprop} &$ prop)
-  (K (ALLGOALS (Simplifier.simp_tac simpset))); (*FIXME*)
-
-in
-
-fun simp_div ctxt ct_n ct_m =
-  let
-    val m = HOLogic.dest_nat (Thm.term_of ct_m);
-  in if m = 0 then Drule.instantiate'[] [SOME ct_n] @{thm div_cert2} else
-    let
-      val thy = ProofContext.theory_of ctxt;
-      val n = HOLogic.dest_nat (Thm.term_of ct_n);
-      val c_n = code_of_val n;
-      val c_m = code_of_val m;
-      val (c_q, c_r) = @{code divmod} c_n c_m;
-      val t_q = term_of_code c_q;
-      val t_r = term_of_code c_r;
-      val ct_q = Thm.cterm_of thy t_q;
-      val ct_r = Thm.cterm_of thy t_r;
-      val thm_r = prove (@{cterm "op < \<Colon> nat \<Rightarrow> _"} &$ ct_r &$ ct_m);
-      val thm_m = prove (@{cterm "(op < \<Colon> nat \<Rightarrow> _) 0"} &$ ct_m);
-      val thm_n = prove (@{cterm "(op = \<Colon> nat \<Rightarrow> _)"} &$ ct_n
-        &$ (@{cterm "(op + \<Colon> nat \<Rightarrow> _)"}
-          &$ (@{cterm "(op * \<Colon> nat \<Rightarrow> _)"} &$ ct_m &$ ct_q) &$ ct_r));
-    in @{thm div_cert1} OF [thm_r, thm_m, thm_n] end
-  end;
-
-end;
-*}
-
-ML_val {*
-  simp_div @{context}
-    @{cterm "Suc (Suc (Suc (Suc (Suc 0))))"}
-    @{cterm "(Suc (Suc 0))"}
-*}
-
-ML_val {*
-  simp_div @{context}
-    (Thm.cterm_of @{theory} (HOLogic.mk_nat 170))
-    (Thm.cterm_of @{theory} (HOLogic.mk_nat 42))
-*}
-
-end
\ No newline at end of file
--- a/src/HOL/ex/Formal_Power_Series_Examples.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/ex/Formal_Power_Series_Examples.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -191,31 +191,39 @@
 lemma E_minus_ii_sin_cos: "E (- (ii * c)) = fps_cos c - fps_const ii * fps_sin c "
   unfolding minus_mult_right Eii_sin_cos by simp
 
+lemma fps_const_minus: "fps_const (c::'a::group_add) - fps_const d = fps_const (c - d)  "by (simp add: fps_eq_iff fps_const_def)
+
+lemma fps_number_of_fps_const: "number_of i = fps_const (number_of i :: 'a:: {comm_ring_1, number_ring})"
+  apply (subst (2) number_of_eq)
+apply(rule int_induct[of _ 0])
+apply (simp_all add: number_of_fps_def)
+by (simp_all add: fps_const_add[symmetric] fps_const_minus[symmetric])
+
 lemma fps_cos_Eii:
   "fps_cos c = (E (ii * c) + E (- ii * c)) / fps_const 2"
 proof-
   have th: "fps_cos c + fps_cos c = fps_cos c * fps_const 2" 
-    by (simp add: fps_eq_iff)
+    by (simp add: fps_eq_iff fps_number_of_fps_const complex_number_of_def[symmetric])
   show ?thesis
   unfolding Eii_sin_cos minus_mult_commute
-  by (simp add: fps_divide_def fps_const_inverse th)
+  by (simp add: fps_number_of_fps_const fps_divide_def fps_const_inverse th complex_number_of_def[symmetric])
 qed
 
 lemma fps_sin_Eii:
   "fps_sin c = (E (ii * c) - E (- ii * c)) / fps_const (2*ii)"
 proof-
   have th: "fps_const \<i> * fps_sin c + fps_const \<i> * fps_sin c = fps_sin c * fps_const (2 * ii)" 
-    by (simp add: fps_eq_iff)
+    by (simp add: fps_eq_iff fps_number_of_fps_const complex_number_of_def[symmetric])
   show ?thesis
   unfolding Eii_sin_cos minus_mult_commute
   by (simp add: fps_divide_def fps_const_inverse th)
 qed
 
 lemma fps_const_mult_2: "fps_const (2::'a::number_ring) * a = a +a"
-  by (simp add: fps_eq_iff)
+  by (simp add: fps_eq_iff fps_number_of_fps_const)
 
 lemma fps_const_mult_2_right: "a * fps_const (2::'a::number_ring) = a +a"
-  by (simp add: fps_eq_iff)
+  by (simp add: fps_eq_iff fps_number_of_fps_const)
 
 lemma fps_tan_Eii:
   "fps_tan c = (E (ii * c) - E (- ii * c)) / (fps_const ii * (E (ii * c) + E (- ii * c)))"
--- a/src/HOL/ex/ImperativeQuicksort.thy	Sat Mar 28 00:11:02 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,637 +0,0 @@
-theory ImperativeQuicksort
-imports "~~/src/HOL/Imperative_HOL/Imperative_HOL" Subarray Multiset Efficient_Nat
-begin
-
-text {* We prove QuickSort correct in the Relational Calculus. *}
-
-definition swap :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap"
-where
-  "swap arr i j = (
-     do
-       x \<leftarrow> nth arr i;
-       y \<leftarrow> nth arr j;
-       upd i y arr;
-       upd j x arr;
-       return ()
-     done)"
-
-lemma swap_permutes:
-  assumes "crel (swap a i j) h h' rs"
-  shows "multiset_of (get_array a h') 
-  = multiset_of (get_array a h)"
-  using assms
-  unfolding swap_def
-  by (auto simp add: Heap.length_def multiset_of_swap dest: sym [of _ "h'"] elim!: crelE crel_nth crel_return crel_upd)
-
-function part1 :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat Heap"
-where
-  "part1 a left right p = (
-     if (right \<le> left) then return right
-     else (do
-       v \<leftarrow> nth a left;
-       (if (v \<le> p) then (part1 a (left + 1) right p)
-                    else (do swap a left right;
-  part1 a left (right - 1) p done))
-     done))"
-by pat_completeness auto
-
-termination
-by (relation "measure (\<lambda>(_,l,r,_). r - l )") auto
-
-declare part1.simps[simp del]
-
-lemma part_permutes:
-  assumes "crel (part1 a l r p) h h' rs"
-  shows "multiset_of (get_array a h') 
-  = multiset_of (get_array a h)"
-  using assms
-proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
-  case (1 a l r p h h' rs)
-  thus ?case
-    unfolding part1.simps [of a l r p]
-    by (elim crelE crel_if crel_return crel_nth) (auto simp add: swap_permutes)
-qed
-
-lemma part_returns_index_in_bounds:
-  assumes "crel (part1 a l r p) h h' rs"
-  assumes "l \<le> r"
-  shows "l \<le> rs \<and> rs \<le> r"
-using assms
-proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
-  case (1 a l r p h h' rs)
-  note cr = `crel (part1 a l r p) h h' rs`
-  show ?case
-  proof (cases "r \<le> l")
-    case True (* Terminating case *)
-    with cr `l \<le> r` show ?thesis
-      unfolding part1.simps[of a l r p]
-      by (elim crelE crel_if crel_return crel_nth) auto
-  next
-    case False (* recursive case *)
-    note rec_condition = this
-    let ?v = "get_array a h ! l"
-    show ?thesis
-    proof (cases "?v \<le> p")
-      case True
-      with cr False
-      have rec1: "crel (part1 a (l + 1) r p) h h' rs"
-        unfolding part1.simps[of a l r p]
-        by (elim crelE crel_nth crel_if crel_return) auto
-      from rec_condition have "l + 1 \<le> r" by arith
-      from 1(1)[OF rec_condition True rec1 `l + 1 \<le> r`]
-      show ?thesis by simp
-    next
-      case False
-      with rec_condition cr
-      obtain h1 where swp: "crel (swap a l r) h h1 ()"
-        and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
-        unfolding part1.simps[of a l r p]
-        by (elim crelE crel_nth crel_if crel_return) auto
-      from rec_condition have "l \<le> r - 1" by arith
-      from 1(2) [OF rec_condition False rec2 `l \<le> r - 1`] show ?thesis by fastsimp
-    qed
-  qed
-qed
-
-lemma part_length_remains:
-  assumes "crel (part1 a l r p) h h' rs"
-  shows "Heap.length a h = Heap.length a h'"
-using assms
-proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
-  case (1 a l r p h h' rs)
-  note cr = `crel (part1 a l r p) h h' rs`
-  
-  show ?case
-  proof (cases "r \<le> l")
-    case True (* Terminating case *)
-    with cr show ?thesis
-      unfolding part1.simps[of a l r p]
-      by (elim crelE crel_if crel_return crel_nth) auto
-  next
-    case False (* recursive case *)
-    with cr 1 show ?thesis
-      unfolding part1.simps [of a l r p] swap_def
-      by (auto elim!: crelE crel_if crel_nth crel_return crel_upd) fastsimp
-  qed
-qed
-
-lemma part_outer_remains:
-  assumes "crel (part1 a l r p) h h' rs"
-  shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
-  using assms
-proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
-  case (1 a l r p h h' rs)
-  note cr = `crel (part1 a l r p) h h' rs`
-  
-  show ?case
-  proof (cases "r \<le> l")
-    case True (* Terminating case *)
-    with cr show ?thesis
-      unfolding part1.simps[of a l r p]
-      by (elim crelE crel_if crel_return crel_nth) auto
-  next
-    case False (* recursive case *)
-    note rec_condition = this
-    let ?v = "get_array a h ! l"
-    show ?thesis
-    proof (cases "?v \<le> p")
-      case True
-      with cr False
-      have rec1: "crel (part1 a (l + 1) r p) h h' rs"
-        unfolding part1.simps[of a l r p]
-        by (elim crelE crel_nth crel_if crel_return) auto
-      from 1(1)[OF rec_condition True rec1]
-      show ?thesis by fastsimp
-    next
-      case False
-      with rec_condition cr
-      obtain h1 where swp: "crel (swap a l r) h h1 ()"
-        and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
-        unfolding part1.simps[of a l r p]
-        by (elim crelE crel_nth crel_if crel_return) auto
-      from swp rec_condition have
-        "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array a h ! i = get_array a h1 ! i"
-	unfolding swap_def
-	by (elim crelE crel_nth crel_upd crel_return) auto
-      with 1(2) [OF rec_condition False rec2] show ?thesis by fastsimp
-    qed
-  qed
-qed
-
-
-lemma part_partitions:
-  assumes "crel (part1 a l r p) h h' rs"
-  shows "(\<forall>i. l \<le> i \<and> i < rs \<longrightarrow> get_array (a::nat array) h' ! i \<le> p)
-  \<and> (\<forall>i. rs < i \<and> i \<le> r \<longrightarrow> get_array a h' ! i \<ge> p)"
-  using assms
-proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
-  case (1 a l r p h h' rs)
-  note cr = `crel (part1 a l r p) h h' rs`
-  
-  show ?case
-  proof (cases "r \<le> l")
-    case True (* Terminating case *)
-    with cr have "rs = r"
-      unfolding part1.simps[of a l r p]
-      by (elim crelE crel_if crel_return crel_nth) auto
-    with True
-    show ?thesis by auto
-  next
-    case False (* recursive case *)
-    note lr = this
-    let ?v = "get_array a h ! l"
-    show ?thesis
-    proof (cases "?v \<le> p")
-      case True
-      with lr cr
-      have rec1: "crel (part1 a (l + 1) r p) h h' rs"
-        unfolding part1.simps[of a l r p]
-        by (elim crelE crel_nth crel_if crel_return) auto
-      from True part_outer_remains[OF rec1] have a_l: "get_array a h' ! l \<le> p"
-	by fastsimp
-      have "\<forall>i. (l \<le> i = (l = i \<or> Suc l \<le> i))" by arith
-      with 1(1)[OF False True rec1] a_l show ?thesis
-	by auto
-    next
-      case False
-      with lr cr
-      obtain h1 where swp: "crel (swap a l r) h h1 ()"
-        and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
-        unfolding part1.simps[of a l r p]
-        by (elim crelE crel_nth crel_if crel_return) auto
-      from swp False have "get_array a h1 ! r \<ge> p"
-	unfolding swap_def
-	by (auto simp add: Heap.length_def elim!: crelE crel_nth crel_upd crel_return)
-      with part_outer_remains [OF rec2] lr have a_r: "get_array a h' ! r \<ge> p"
-	by fastsimp
-      have "\<forall>i. (i \<le> r = (i = r \<or> i \<le> r - 1))" by arith
-      with 1(2)[OF lr False rec2] a_r show ?thesis
-	by auto
-    qed
-  qed
-qed
-
-
-fun partition :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat Heap"
-where
-  "partition a left right = (do
-     pivot \<leftarrow> nth a right;
-     middle \<leftarrow> part1 a left (right - 1) pivot;
-     v \<leftarrow> nth a middle;
-     m \<leftarrow> return (if (v \<le> pivot) then (middle + 1) else middle);
-     swap a m right;
-     return m
-   done)"
-
-declare partition.simps[simp del]
-
-lemma partition_permutes:
-  assumes "crel (partition a l r) h h' rs"
-  shows "multiset_of (get_array a h') 
-  = multiset_of (get_array a h)"
-proof -
-    from assms part_permutes swap_permutes show ?thesis
-      unfolding partition.simps
-      by (elim crelE crel_return crel_nth crel_if crel_upd) auto
-qed
-
-lemma partition_length_remains:
-  assumes "crel (partition a l r) h h' rs"
-  shows "Heap.length a h = Heap.length a h'"
-proof -
-  from assms part_length_remains show ?thesis
-    unfolding partition.simps swap_def
-    by (elim crelE crel_return crel_nth crel_if crel_upd) auto
-qed
-
-lemma partition_outer_remains:
-  assumes "crel (partition a l r) h h' rs"
-  assumes "l < r"
-  shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
-proof -
-  from assms part_outer_remains part_returns_index_in_bounds show ?thesis
-    unfolding partition.simps swap_def
-    by (elim crelE crel_return crel_nth crel_if crel_upd) fastsimp
-qed
-
-lemma partition_returns_index_in_bounds:
-  assumes crel: "crel (partition a l r) h h' rs"
-  assumes "l < r"
-  shows "l \<le> rs \<and> rs \<le> r"
-proof -
-  from crel obtain middle h'' p where part: "crel (part1 a l (r - 1) p) h h'' middle"
-    and rs_equals: "rs = (if get_array a h'' ! middle \<le> get_array a h ! r then middle + 1
-         else middle)"
-    unfolding partition.simps
-    by (elim crelE crel_return crel_nth crel_if crel_upd) simp
-  from `l < r` have "l \<le> r - 1" by arith
-  from part_returns_index_in_bounds[OF part this] rs_equals `l < r` show ?thesis by auto
-qed
-
-lemma partition_partitions:
-  assumes crel: "crel (partition a l r) h h' rs"
-  assumes "l < r"
-  shows "(\<forall>i. l \<le> i \<and> i < rs \<longrightarrow> get_array (a::nat array) h' ! i \<le> get_array a h' ! rs) \<and>
-  (\<forall>i. rs < i \<and> i \<le> r \<longrightarrow> get_array a h' ! rs \<le> get_array a h' ! i)"
-proof -
-  let ?pivot = "get_array a h ! r" 
-  from crel obtain middle h1 where part: "crel (part1 a l (r - 1) ?pivot) h h1 middle"
-    and swap: "crel (swap a rs r) h1 h' ()"
-    and rs_equals: "rs = (if get_array a h1 ! middle \<le> ?pivot then middle + 1
-         else middle)"
-    unfolding partition.simps
-    by (elim crelE crel_return crel_nth crel_if crel_upd) simp
-  from swap have h'_def: "h' = Heap.upd a r (get_array a h1 ! rs)
-    (Heap.upd a rs (get_array a h1 ! r) h1)"
-    unfolding swap_def
-    by (elim crelE crel_return crel_nth crel_upd) simp
-  from swap have in_bounds: "r < Heap.length a h1 \<and> rs < Heap.length a h1"
-    unfolding swap_def
-    by (elim crelE crel_return crel_nth crel_upd) simp
-  from swap have swap_length_remains: "Heap.length a h1 = Heap.length a h'"
-    unfolding swap_def by (elim crelE crel_return crel_nth crel_upd) auto
-  from `l < r` have "l \<le> r - 1" by simp 
-  note middle_in_bounds = part_returns_index_in_bounds[OF part this]
-  from part_outer_remains[OF part] `l < r`
-  have "get_array a h ! r = get_array a h1 ! r"
-    by fastsimp
-  with swap
-  have right_remains: "get_array a h ! r = get_array a h' ! rs"
-    unfolding swap_def
-    by (auto simp add: Heap.length_def elim!: crelE crel_return crel_nth crel_upd) (cases "r = rs", auto)
-  from part_partitions [OF part]
-  show ?thesis
-  proof (cases "get_array a h1 ! middle \<le> ?pivot")
-    case True
-    with rs_equals have rs_equals: "rs = middle + 1" by simp
-    { 
-      fix i
-      assume i_is_left: "l \<le> i \<and> i < rs"
-      with swap_length_remains in_bounds middle_in_bounds rs_equals `l < r`
-      have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
-      from i_is_left rs_equals have "l \<le> i \<and> i < middle \<or> i = middle" by arith
-      with part_partitions[OF part] right_remains True
-      have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
-      with i_props h'_def in_bounds have "get_array a h' ! i \<le> get_array a h' ! rs"
-	unfolding Heap.upd_def Heap.length_def by simp
-    }
-    moreover
-    {
-      fix i
-      assume "rs < i \<and> i \<le> r"
-
-      hence "(rs < i \<and> i \<le> r - 1) \<or> (rs < i \<and> i = r)" by arith
-      hence "get_array a h' ! rs \<le> get_array a h' ! i"
-      proof
-	assume i_is: "rs < i \<and> i \<le> r - 1"
-	with swap_length_remains in_bounds middle_in_bounds rs_equals
-	have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
-	from part_partitions[OF part] rs_equals right_remains i_is
-	have "get_array a h' ! rs \<le> get_array a h1 ! i"
-	  by fastsimp
-	with i_props h'_def show ?thesis by fastsimp
-      next
-	assume i_is: "rs < i \<and> i = r"
-	with rs_equals have "Suc middle \<noteq> r" by arith
-	with middle_in_bounds `l < r` have "Suc middle \<le> r - 1" by arith
-	with part_partitions[OF part] right_remains 
-	have "get_array a h' ! rs \<le> get_array a h1 ! (Suc middle)"
-	  by fastsimp
-	with i_is True rs_equals right_remains h'_def
-	show ?thesis using in_bounds
-	  unfolding Heap.upd_def Heap.length_def
-	  by auto
-      qed
-    }
-    ultimately show ?thesis by auto
-  next
-    case False
-    with rs_equals have rs_equals: "middle = rs" by simp
-    { 
-      fix i
-      assume i_is_left: "l \<le> i \<and> i < rs"
-      with swap_length_remains in_bounds middle_in_bounds rs_equals
-      have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
-      from part_partitions[OF part] rs_equals right_remains i_is_left
-      have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
-      with i_props h'_def have "get_array a h' ! i \<le> get_array a h' ! rs"
-	unfolding Heap.upd_def by simp
-    }
-    moreover
-    {
-      fix i
-      assume "rs < i \<and> i \<le> r"
-      hence "(rs < i \<and> i \<le> r - 1) \<or> i = r" by arith
-      hence "get_array a h' ! rs \<le> get_array a h' ! i"
-      proof
-	assume i_is: "rs < i \<and> i \<le> r - 1"
-	with swap_length_remains in_bounds middle_in_bounds rs_equals
-	have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
-	from part_partitions[OF part] rs_equals right_remains i_is
-	have "get_array a h' ! rs \<le> get_array a h1 ! i"
-	  by fastsimp
-	with i_props h'_def show ?thesis by fastsimp
-      next
-	assume i_is: "i = r"
-	from i_is False rs_equals right_remains h'_def
-	show ?thesis using in_bounds
-	  unfolding Heap.upd_def Heap.length_def
-	  by auto
-      qed
-    }
-    ultimately
-    show ?thesis by auto
-  qed
-qed
-
-
-function quicksort :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap"
-where
-  "quicksort arr left right =
-     (if (right > left)  then
-        do
-          pivotNewIndex \<leftarrow> partition arr left right;
-          pivotNewIndex \<leftarrow> assert (\<lambda>x. left \<le> x \<and> x \<le> right) pivotNewIndex;
-          quicksort arr left (pivotNewIndex - 1);
-          quicksort arr (pivotNewIndex + 1) right
-        done
-     else return ())"
-by pat_completeness auto
-
-(* For termination, we must show that the pivotNewIndex is between left and right *) 
-termination
-by (relation "measure (\<lambda>(a, l, r). (r - l))") auto
-
-declare quicksort.simps[simp del]
-
-
-lemma quicksort_permutes:
-  assumes "crel (quicksort a l r) h h' rs"
-  shows "multiset_of (get_array a h') 
-  = multiset_of (get_array a h)"
-  using assms
-proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
-  case (1 a l r h h' rs)
-  with partition_permutes show ?case
-    unfolding quicksort.simps [of a l r]
-    by (elim crel_if crelE crel_assert crel_return) auto
-qed
-
-lemma length_remains:
-  assumes "crel (quicksort a l r) h h' rs"
-  shows "Heap.length a h = Heap.length a h'"
-using assms
-proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
-  case (1 a l r h h' rs)
-  with partition_length_remains show ?case
-    unfolding quicksort.simps [of a l r]
-    by (elim crel_if crelE crel_assert crel_return) auto
-qed
-
-lemma quicksort_outer_remains:
-  assumes "crel (quicksort a l r) h h' rs"
-   shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
-  using assms
-proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
-  case (1 a l r h h' rs)
-  note cr = `crel (quicksort a l r) h h' rs`
-  thus ?case
-  proof (cases "r > l")
-    case False
-    with cr have "h' = h"
-      unfolding quicksort.simps [of a l r]
-      by (elim crel_if crel_return) auto
-    thus ?thesis by simp
-  next
-  case True
-   { 
-      fix h1 h2 p ret1 ret2 i
-      assume part: "crel (partition a l r) h h1 p"
-      assume qs1: "crel (quicksort a l (p - 1)) h1 h2 ret1"
-      assume qs2: "crel (quicksort a (p + 1) r) h2 h' ret2"
-      assume pivot: "l \<le> p \<and> p \<le> r"
-      assume i_outer: "i < l \<or> r < i"
-      from  partition_outer_remains [OF part True] i_outer
-      have "get_array a h !i = get_array a h1 ! i" by fastsimp
-      moreover
-      with 1(1) [OF True pivot qs1] pivot i_outer
-      have "get_array a h1 ! i = get_array a h2 ! i" by auto
-      moreover
-      with qs2 1(2) [of p h2 h' ret2] True pivot i_outer
-      have "get_array a h2 ! i = get_array a h' ! i" by auto
-      ultimately have "get_array a h ! i= get_array a h' ! i" by simp
-    }
-    with cr show ?thesis
-      unfolding quicksort.simps [of a l r]
-      by (elim crel_if crelE crel_assert crel_return) auto
-  qed
-qed
-
-lemma quicksort_is_skip:
-  assumes "crel (quicksort a l r) h h' rs"
-  shows "r \<le> l \<longrightarrow> h = h'"
-  using assms
-  unfolding quicksort.simps [of a l r]
-  by (elim crel_if crel_return) auto
- 
-lemma quicksort_sorts:
-  assumes "crel (quicksort a l r) h h' rs"
-  assumes l_r_length: "l < Heap.length a h" "r < Heap.length a h" 
-  shows "sorted (subarray l (r + 1) a h')"
-  using assms
-proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
-  case (1 a l r h h' rs)
-  note cr = `crel (quicksort a l r) h h' rs`
-  thus ?case
-  proof (cases "r > l")
-    case False
-    hence "l \<ge> r + 1 \<or> l = r" by arith 
-    with length_remains[OF cr] 1(5) show ?thesis
-      by (auto simp add: subarray_Nil subarray_single)
-  next
-    case True
-    { 
-      fix h1 h2 p
-      assume part: "crel (partition a l r) h h1 p"
-      assume qs1: "crel (quicksort a l (p - 1)) h1 h2 ()"
-      assume qs2: "crel (quicksort a (p + 1) r) h2 h' ()"
-      from partition_returns_index_in_bounds [OF part True]
-      have pivot: "l\<le> p \<and> p \<le> r" .
-     note length_remains = length_remains[OF qs2] length_remains[OF qs1] partition_length_remains[OF part]
-      from quicksort_outer_remains [OF qs2] quicksort_outer_remains [OF qs1] pivot quicksort_is_skip[OF qs1]
-      have pivot_unchanged: "get_array a h1 ! p = get_array a h' ! p" by (cases p, auto)
-        (*-- First of all, by induction hypothesis both sublists are sorted. *)
-      from 1(1)[OF True pivot qs1] length_remains pivot 1(5) 
-      have IH1: "sorted (subarray l p a h2)"  by (cases p, auto simp add: subarray_Nil)
-      from quicksort_outer_remains [OF qs2] length_remains
-      have left_subarray_remains: "subarray l p a h2 = subarray l p a h'"
-	by (simp add: subarray_eq_samelength_iff)
-      with IH1 have IH1': "sorted (subarray l p a h')" by simp
-      from 1(2)[OF True pivot qs2] pivot 1(5) length_remains
-      have IH2: "sorted (subarray (p + 1) (r + 1) a h')"
-        by (cases "Suc p \<le> r", auto simp add: subarray_Nil)
-           (* -- Secondly, both sublists remain partitioned. *)
-      from partition_partitions[OF part True]
-      have part_conds1: "\<forall>j. j \<in> set (subarray l p a h1) \<longrightarrow> j \<le> get_array a h1 ! p "
-        and part_conds2: "\<forall>j. j \<in> set (subarray (p + 1) (r + 1) a h1) \<longrightarrow> get_array a h1 ! p \<le> j"
-        by (auto simp add: all_in_set_subarray_conv)
-      from quicksort_outer_remains [OF qs1] quicksort_permutes [OF qs1] True
-        length_remains 1(5) pivot multiset_of_sublist [of l p "get_array a h1" "get_array a h2"]
-      have multiset_partconds1: "multiset_of (subarray l p a h2) = multiset_of (subarray l p a h1)"
-	unfolding Heap.length_def subarray_def by (cases p, auto)
-      with left_subarray_remains part_conds1 pivot_unchanged
-      have part_conds2': "\<forall>j. j \<in> set (subarray l p a h') \<longrightarrow> j \<le> get_array a h' ! p"
-        by (simp, subst set_of_multiset_of[symmetric], simp)
-          (* -- These steps are the analogous for the right sublist \<dots> *)
-      from quicksort_outer_remains [OF qs1] length_remains
-      have right_subarray_remains: "subarray (p + 1) (r + 1) a h1 = subarray (p + 1) (r + 1) a h2"
-	by (auto simp add: subarray_eq_samelength_iff)
-      from quicksort_outer_remains [OF qs2] quicksort_permutes [OF qs2] True
-        length_remains 1(5) pivot multiset_of_sublist [of "p + 1" "r + 1" "get_array a h2" "get_array a h'"]
-      have multiset_partconds2: "multiset_of (subarray (p + 1) (r + 1) a h') = multiset_of (subarray (p + 1) (r + 1) a h2)"
-        unfolding Heap.length_def subarray_def by auto
-      with right_subarray_remains part_conds2 pivot_unchanged
-      have part_conds1': "\<forall>j. j \<in> set (subarray (p + 1) (r + 1) a h') \<longrightarrow> get_array a h' ! p \<le> j"
-        by (simp, subst set_of_multiset_of[symmetric], simp)
-          (* -- Thirdly and finally, we show that the array is sorted
-          following from the facts above. *)
-      from True pivot 1(5) length_remains have "subarray l (r + 1) a h' = subarray l p a h' @ [get_array a h' ! p] @ subarray (p + 1) (r + 1) a h'"
-	by (simp add: subarray_nth_array_Cons, cases "l < p") (auto simp add: subarray_append subarray_Nil)
-      with IH1' IH2 part_conds1' part_conds2' pivot have ?thesis
-	unfolding subarray_def
-	apply (auto simp add: sorted_append sorted_Cons all_in_set_sublist'_conv)
-	by (auto simp add: set_sublist' dest: le_trans [of _ "get_array a h' ! p"])
-    }
-    with True cr show ?thesis
-      unfolding quicksort.simps [of a l r]
-      by (elim crel_if crel_return crelE crel_assert) auto
-  qed
-qed
-
-
-lemma quicksort_is_sort:
-  assumes crel: "crel (quicksort a 0 (Heap.length a h - 1)) h h' rs"
-  shows "get_array a h' = sort (get_array a h)"
-proof (cases "get_array a h = []")
-  case True
-  with quicksort_is_skip[OF crel] show ?thesis
-  unfolding Heap.length_def by simp
-next
-  case False
-  from quicksort_sorts [OF crel] False have "sorted (sublist' 0 (List.length (get_array a h)) (get_array a h'))"
-    unfolding Heap.length_def subarray_def by auto
-  with length_remains[OF crel] have "sorted (get_array a h')"
-    unfolding Heap.length_def by simp
-  with quicksort_permutes [OF crel] properties_for_sort show ?thesis by fastsimp
-qed
-
-subsection {* No Errors in quicksort *}
-text {* We have proved that quicksort sorts (if no exceptions occur).
-We will now show that exceptions do not occur. *}
-
-lemma noError_part1: 
-  assumes "l < Heap.length a h" "r < Heap.length a h"
-  shows "noError (part1 a l r p) h"
-  using assms
-proof (induct a l r p arbitrary: h rule: part1.induct)
-  case (1 a l r p)
-  thus ?case
-    unfolding part1.simps [of a l r] swap_def
-    by (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd elim!: crelE crel_upd crel_nth crel_return)
-qed
-
-lemma noError_partition:
-  assumes "l < r" "l < Heap.length a h" "r < Heap.length a h"
-  shows "noError (partition a l r) h"
-using assms
-unfolding partition.simps swap_def
-apply (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd noError_part1 elim!: crelE crel_upd crel_nth crel_return)
-apply (frule part_length_remains)
-apply (frule part_returns_index_in_bounds)
-apply auto
-apply (frule part_length_remains)
-apply (frule part_returns_index_in_bounds)
-apply auto
-apply (frule part_length_remains)
-apply auto
-done
-
-lemma noError_quicksort:
-  assumes "l < Heap.length a h" "r < Heap.length a h"
-  shows "noError (quicksort a l r) h"
-using assms
-proof (induct a l r arbitrary: h rule: quicksort.induct)
-  case (1 a l ri h)
-  thus ?case
-    unfolding quicksort.simps [of a l ri]
-    apply (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd noError_assert noError_partition)
-    apply (frule partition_returns_index_in_bounds)
-    apply auto
-    apply (frule partition_returns_index_in_bounds)
-    apply auto
-    apply (auto elim!: crel_assert dest!: partition_length_remains length_remains)
-    apply (subgoal_tac "Suc r \<le> ri \<or> r = ri") 
-    apply (erule disjE)
-    apply auto
-    unfolding quicksort.simps [of a "Suc ri" ri]
-    apply (auto intro!: noError_if noError_return)
-    done
-qed
-
-
-subsection {* Example *}
-
-definition "qsort a = do
-    k \<leftarrow> length a;
-    quicksort a 0 (k - 1);
-    return a
-  done"
-
-ML {* @{code qsort} (Array.fromList [42, 2, 3, 5, 0, 1705, 8, 3, 15]) () *}
-
-export_code qsort in SML_imp module_name QSort
-export_code qsort in OCaml module_name QSort file -
-export_code qsort in OCaml_imp module_name QSort file -
-export_code qsort in Haskell module_name QSort file -
-
-end
\ No newline at end of file
--- a/src/HOL/ex/Lagrange.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/ex/Lagrange.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/ex/Lagrange.thy
-    ID:         $Id$
     Author:     Tobias Nipkow
     Copyright   1996 TU Muenchen
 *)
@@ -35,7 +34,7 @@
    sq (x1*y2 + x2*y1 + x3*y4 - x4*y3)  +
    sq (x1*y3 - x2*y4 + x3*y1 + x4*y2)  +
    sq (x1*y4 + x2*y3 - x3*y2 + x4*y1)"
-by (simp add: sq_def algebra_simps)
+by (simp only: sq_def ring_simps)
 
 
 text {* A challenge by John Harrison. Takes about 12s on a 1.6GHz machine. *}
@@ -51,6 +50,6 @@
       sq (p1*u2 + q1*t2 - r1*w2 + s1*v2 - t1*q2 + u1*p2 - v1*s2 + w1*r2) +
       sq (p1*v2 + q1*w2 + r1*t2 - s1*u2 - t1*r2 + u1*s2 + v1*p2 - w1*q2) +
       sq (p1*w2 - q1*v2 + r1*u2 + s1*t2 - t1*s2 - u1*r2 + v1*q2 + w1*p2)"
-by (simp add: sq_def algebra_simps)
+by (simp only: sq_def ring_simps)
 
 end
--- a/src/HOL/ex/LocaleTest2.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/ex/LocaleTest2.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -468,7 +468,7 @@
 
 subsubsection {* Total order @{text "<="} on @{typ int} *}
 
-interpretation int!: dpo "op <= :: [int, int] => bool"
+interpretation int: dpo "op <= :: [int, int] => bool"
   where "(dpo.less (op <=) (x::int) y) = (x < y)"
   txt {* We give interpretation for less, but not @{text is_inf} and @{text is_sub}. *}
 proof -
@@ -487,7 +487,7 @@
 lemma "(op < :: [int, int] => bool) = op <"
   apply (rule int.abs_test) done
 
-interpretation int!: dlat "op <= :: [int, int] => bool"
+interpretation int: dlat "op <= :: [int, int] => bool"
   where meet_eq: "dlat.meet (op <=) (x::int) y = min x y"
     and join_eq: "dlat.join (op <=) (x::int) y = max x y"
 proof -
@@ -511,7 +511,7 @@
     by auto
 qed
 
-interpretation int!: dlo "op <= :: [int, int] => bool"
+interpretation int: dlo "op <= :: [int, int] => bool"
   proof qed arith
 
 text {* Interpreted theorems from the locales, involving defined terms. *}
@@ -524,7 +524,7 @@
 
 subsubsection {* Total order @{text "<="} on @{typ nat} *}
 
-interpretation nat!: dpo "op <= :: [nat, nat] => bool"
+interpretation nat: dpo "op <= :: [nat, nat] => bool"
   where "dpo.less (op <=) (x::nat) y = (x < y)"
   txt {* We give interpretation for less, but not @{text is_inf} and @{text is_sub}. *}
 proof -
@@ -538,7 +538,7 @@
     done
 qed
 
-interpretation nat!: dlat "op <= :: [nat, nat] => bool"
+interpretation nat: dlat "op <= :: [nat, nat] => bool"
   where "dlat.meet (op <=) (x::nat) y = min x y"
     and "dlat.join (op <=) (x::nat) y = max x y"
 proof -
@@ -562,7 +562,7 @@
     by auto
 qed
 
-interpretation nat!: dlo "op <= :: [nat, nat] => bool"
+interpretation nat: dlo "op <= :: [nat, nat] => bool"
   proof qed arith
 
 text {* Interpreted theorems from the locales, involving defined terms. *}
@@ -575,7 +575,7 @@
 
 subsubsection {* Lattice @{text "dvd"} on @{typ nat} *}
 
-interpretation nat_dvd!: dpo "op dvd :: [nat, nat] => bool"
+interpretation nat_dvd: dpo "op dvd :: [nat, nat] => bool"
   where "dpo.less (op dvd) (x::nat) y = (x dvd y & x ~= y)"
   txt {* We give interpretation for less, but not @{text is_inf} and @{text is_sub}. *}
 proof -
@@ -589,7 +589,7 @@
     done
 qed
 
-interpretation nat_dvd!: dlat "op dvd :: [nat, nat] => bool"
+interpretation nat_dvd: dlat "op dvd :: [nat, nat] => bool"
   where "dlat.meet (op dvd) (x::nat) y = gcd x y"
     and "dlat.join (op dvd) (x::nat) y = lcm x y"
 proof -
@@ -834,7 +834,7 @@
 
 subsubsection {* Interpretation of Functions *}
 
-interpretation Dfun!: Dmonoid "op o" "id :: 'a => 'a"
+interpretation Dfun: Dmonoid "op o" "id :: 'a => 'a"
   where "Dmonoid.unit (op o) id f = bij (f::'a => 'a)"
 (*    and "Dmonoid.inv (op o) id" = "inv :: ('a => 'a) => ('a => 'a)" *)
 proof -
@@ -884,7 +884,7 @@
   "(f :: unit => unit) = id"
   by rule simp
 
-interpretation Dfun!: Dgrp "op o" "id :: unit => unit"
+interpretation Dfun: Dgrp "op o" "id :: unit => unit"
   where "Dmonoid.inv (op o) id f = inv (f :: unit => unit)"
 proof -
   have "Dmonoid op o (id :: 'a => 'a)" ..
--- a/src/HOL/ex/MergeSort.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/ex/MergeSort.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/ex/Merge.thy
-    ID:         $Id$
     Author:     Tobias Nipkow
     Copyright   2002 TU Muenchen
 *)
--- a/src/HOL/ex/ROOT.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOL/ex/ROOT.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -21,7 +21,6 @@
 
 use_thys [
   "Numeral",
-  "ImperativeQuicksort",
   "Higher_Order_Logic",
   "Abstract_NAT",
   "Guess",
@@ -57,7 +56,6 @@
   "Classical",
   "set",
   "Meson_Test",
-  "Code_Antiq",
   "Termination",
   "Coherent",
   "PresburgerEx",
--- a/src/HOL/ex/Subarray.thy	Sat Mar 28 00:11:02 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,66 +0,0 @@
-theory Subarray
-imports Array Sublist
-begin
-
-definition subarray :: "nat \<Rightarrow> nat \<Rightarrow> ('a::heap) array \<Rightarrow> heap \<Rightarrow> 'a list"
-where
-  "subarray n m a h \<equiv> sublist' n m (get_array a h)"
-
-lemma subarray_upd: "i \<ge> m \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
-apply (simp add: subarray_def Heap.upd_def)
-apply (simp add: sublist'_update1)
-done
-
-lemma subarray_upd2: " i < n  \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
-apply (simp add: subarray_def Heap.upd_def)
-apply (subst sublist'_update2)
-apply fastsimp
-apply simp
-done
-
-lemma subarray_upd3: "\<lbrakk> n \<le> i; i < m\<rbrakk> \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h[i - n := v]"
-unfolding subarray_def Heap.upd_def
-by (simp add: sublist'_update3)
-
-lemma subarray_Nil: "n \<ge> m \<Longrightarrow> subarray n m a h = []"
-by (simp add: subarray_def sublist'_Nil')
-
-lemma subarray_single: "\<lbrakk> n < Heap.length a h \<rbrakk> \<Longrightarrow> subarray n (Suc n) a h = [get_array a h ! n]" 
-by (simp add: subarray_def Heap.length_def sublist'_single)
-
-lemma length_subarray: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray n m a h) = m - n"
-by (simp add: subarray_def Heap.length_def length_sublist')
-
-lemma length_subarray_0: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray 0 m a h) = m"
-by (simp add: length_subarray)
-
-lemma subarray_nth_array_Cons: "\<lbrakk> i < Heap.length a h; i < j \<rbrakk> \<Longrightarrow> (get_array a h ! i) # subarray (Suc i) j a h = subarray i j a h"
-unfolding Heap.length_def subarray_def
-by (simp add: sublist'_front)
-
-lemma subarray_nth_array_back: "\<lbrakk> i < j; j \<le> Heap.length a h\<rbrakk> \<Longrightarrow> subarray i j a h = subarray i (j - 1) a h @ [get_array a h ! (j - 1)]"
-unfolding Heap.length_def subarray_def
-by (simp add: sublist'_back)
-
-lemma subarray_append: "\<lbrakk> i < j; j < k \<rbrakk> \<Longrightarrow> subarray i j a h @ subarray j k a h = subarray i k a h"
-unfolding subarray_def
-by (simp add: sublist'_append)
-
-lemma subarray_all: "subarray 0 (Heap.length a h) a h = get_array a h"
-unfolding Heap.length_def subarray_def
-by (simp add: sublist'_all)
-
-lemma nth_subarray: "\<lbrakk> k < j - i; j \<le> Heap.length a h \<rbrakk> \<Longrightarrow> subarray i j a h ! k = get_array a h ! (i + k)"
-unfolding Heap.length_def subarray_def
-by (simp add: nth_sublist')
-
-lemma subarray_eq_samelength_iff: "Heap.length a h = Heap.length a h' \<Longrightarrow> (subarray i j a h = subarray i j a h') = (\<forall>i'. i \<le> i' \<and> i' < j \<longrightarrow> get_array a h ! i' = get_array a h' ! i')"
-unfolding Heap.length_def subarray_def by (rule sublist'_eq_samelength_iff)
-
-lemma all_in_set_subarray_conv: "(\<forall>j. j \<in> set (subarray l r a h) \<longrightarrow> P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
-unfolding subarray_def Heap.length_def by (rule all_in_set_sublist'_conv)
-
-lemma ball_in_set_subarray_conv: "(\<forall>j \<in> set (subarray l r a h). P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
-unfolding subarray_def Heap.length_def by (rule ball_in_set_sublist'_conv)
-
-end
\ No newline at end of file
--- a/src/HOL/ex/Sublist.thy	Sat Mar 28 00:11:02 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,505 +0,0 @@
-(* $Id$ *)
-
-header {* Slices of lists *}
-
-theory Sublist
-imports Multiset
-begin
-
-
-lemma sublist_split: "i \<le> j \<and> j \<le> k \<Longrightarrow> sublist xs {i..<j} @ sublist xs {j..<k} = sublist xs {i..<k}" 
-apply (induct xs arbitrary: i j k)
-apply simp
-apply (simp only: sublist_Cons)
-apply simp
-apply safe
-apply simp
-apply (erule_tac x="0" in meta_allE)
-apply (erule_tac x="j - 1" in meta_allE)
-apply (erule_tac x="k - 1" in meta_allE)
-apply (subgoal_tac "0 \<le> j - 1 \<and> j - 1 \<le> k - 1")
-apply simp
-apply (subgoal_tac "{ja. Suc ja < j} = {0..<j - Suc 0}")
-apply (subgoal_tac "{ja. j \<le> Suc ja \<and> Suc ja < k} = {j - Suc 0..<k - Suc 0}")
-apply (subgoal_tac "{j. Suc j < k} = {0..<k - Suc 0}")
-apply simp
-apply fastsimp
-apply fastsimp
-apply fastsimp
-apply fastsimp
-apply (erule_tac x="i - 1" in meta_allE)
-apply (erule_tac x="j - 1" in meta_allE)
-apply (erule_tac x="k - 1" in meta_allE)
-apply (subgoal_tac " {ja. i \<le> Suc ja \<and> Suc ja < j} = {i - 1 ..<j - 1}")
-apply (subgoal_tac " {ja. j \<le> Suc ja \<and> Suc ja < k} = {j - 1..<k - 1}")
-apply (subgoal_tac "{j. i \<le> Suc j \<and> Suc j < k} = {i - 1..<k - 1}")
-apply (subgoal_tac " i - 1 \<le> j - 1 \<and> j - 1 \<le> k - 1")
-apply simp
-apply fastsimp
-apply fastsimp
-apply fastsimp
-apply fastsimp
-done
-
-lemma sublist_update1: "i \<notin> inds \<Longrightarrow> sublist (xs[i := v]) inds = sublist xs inds"
-apply (induct xs arbitrary: i inds)
-apply simp
-apply (case_tac i)
-apply (simp add: sublist_Cons)
-apply (simp add: sublist_Cons)
-done
-
-lemma sublist_update2: "i \<in> inds \<Longrightarrow> sublist (xs[i := v]) inds = (sublist xs inds)[(card {k \<in> inds. k < i}):= v]"
-proof (induct xs arbitrary: i inds)
-  case Nil thus ?case by simp
-next
-  case (Cons x xs)
-  thus ?case
-  proof (cases i)
-    case 0 with Cons show ?thesis by (simp add: sublist_Cons)
-  next
-    case (Suc i')
-    with Cons show ?thesis
-      apply simp
-      apply (simp add: sublist_Cons)
-      apply auto
-      apply (auto simp add: nat.split)
-      apply (simp add: card_less_Suc[symmetric])
-      apply (simp add: card_less_Suc2)
-      done
-  qed
-qed
-
-lemma sublist_update: "sublist (xs[i := v]) inds = (if i \<in> inds then (sublist xs inds)[(card {k \<in> inds. k < i}) := v] else sublist xs inds)"
-by (simp add: sublist_update1 sublist_update2)
-
-lemma sublist_take: "sublist xs {j. j < m} = take m xs"
-apply (induct xs arbitrary: m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (simp add: sublist_Cons)
-done
-
-lemma sublist_take': "sublist xs {0..<m} = take m xs"
-apply (induct xs arbitrary: m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (simp add: sublist_Cons sublist_take)
-done
-
-lemma sublist_all[simp]: "sublist xs {j. j < length xs} = xs"
-apply (induct xs)
-apply simp
-apply (simp add: sublist_Cons)
-done
-
-lemma sublist_all'[simp]: "sublist xs {0..<length xs} = xs"
-apply (induct xs)
-apply simp
-apply (simp add: sublist_Cons)
-done
-
-lemma sublist_single: "a < length xs \<Longrightarrow> sublist xs {a} = [xs ! a]"
-apply (induct xs arbitrary: a)
-apply simp
-apply(case_tac aa)
-apply simp
-apply (simp add: sublist_Cons)
-apply simp
-apply (simp add: sublist_Cons)
-done
-
-lemma sublist_is_Nil: "\<forall>i \<in> inds. i \<ge> length xs \<Longrightarrow> sublist xs inds = []" 
-apply (induct xs arbitrary: inds)
-apply simp
-apply (simp add: sublist_Cons)
-apply auto
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply auto
-done
-
-lemma sublist_Nil': "sublist xs inds = [] \<Longrightarrow> \<forall>i \<in> inds. i \<ge> length xs"
-apply (induct xs arbitrary: inds)
-apply simp
-apply (simp add: sublist_Cons)
-apply (auto split: if_splits)
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply (case_tac x, auto)
-done
-
-lemma sublist_Nil[simp]: "(sublist xs inds = []) = (\<forall>i \<in> inds. i \<ge> length xs)"
-apply (induct xs arbitrary: inds)
-apply simp
-apply (simp add: sublist_Cons)
-apply auto
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply (case_tac x, auto)
-done
-
-lemma sublist_eq_subseteq: " \<lbrakk> inds' \<subseteq> inds; sublist xs inds = sublist ys inds \<rbrakk> \<Longrightarrow> sublist xs inds' = sublist ys inds'"
-apply (induct xs arbitrary: ys inds inds')
-apply simp
-apply (drule sym, rule sym)
-apply (simp add: sublist_Nil, fastsimp)
-apply (case_tac ys)
-apply (simp add: sublist_Nil, fastsimp)
-apply (auto simp add: sublist_Cons)
-apply (erule_tac x="list" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds'}" in meta_allE)
-apply fastsimp
-apply (erule_tac x="list" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds'}" in meta_allE)
-apply fastsimp
-done
-
-lemma sublist_eq: "\<lbrakk> \<forall>i \<in> inds. ((i < length xs) \<and> (i < length ys)) \<or> ((i \<ge> length xs ) \<and> (i \<ge> length ys)); \<forall>i \<in> inds. xs ! i = ys ! i \<rbrakk> \<Longrightarrow> sublist xs inds = sublist ys inds"
-apply (induct xs arbitrary: ys inds)
-apply simp
-apply (rule sym, simp add: sublist_Nil)
-apply (case_tac ys)
-apply (simp add: sublist_Nil)
-apply (auto simp add: sublist_Cons)
-apply (erule_tac x="list" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply fastsimp
-apply (erule_tac x="list" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply fastsimp
-done
-
-lemma sublist_eq_samelength: "\<lbrakk> length xs = length ys; \<forall>i \<in> inds. xs ! i = ys ! i \<rbrakk> \<Longrightarrow> sublist xs inds = sublist ys inds"
-by (rule sublist_eq, auto)
-
-lemma sublist_eq_samelength_iff: "length xs = length ys \<Longrightarrow> (sublist xs inds = sublist ys inds) = (\<forall>i \<in> inds. xs ! i = ys ! i)"
-apply (induct xs arbitrary: ys inds)
-apply simp
-apply (rule sym, simp add: sublist_Nil)
-apply (case_tac ys)
-apply (simp add: sublist_Nil)
-apply (auto simp add: sublist_Cons)
-apply (case_tac i)
-apply auto
-apply (case_tac i)
-apply auto
-done
-
-section {* Another sublist function *}
-
-function sublist' :: "nat \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where
-  "sublist' n m [] = []"
-| "sublist' n 0 xs = []"
-| "sublist' 0 (Suc m) (x#xs) = (x#sublist' 0 m xs)"
-| "sublist' (Suc n) (Suc m) (x#xs) = sublist' n m xs"
-by pat_completeness auto
-termination by lexicographic_order
-
-subsection {* Proving equivalence to the other sublist command *}
-
-lemma sublist'_sublist: "sublist' n m xs = sublist xs {j. n \<le> j \<and> j < m}"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac n)
-apply (case_tac m)
-apply simp
-apply (simp add: sublist_Cons)
-apply (case_tac m)
-apply simp
-apply (simp add: sublist_Cons)
-done
-
-
-lemma "sublist' n m xs = sublist xs {n..<m}"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac n, case_tac m)
-apply simp
-apply simp
-apply (simp add: sublist_take')
-apply (case_tac m)
-apply simp
-apply (simp add: sublist_Cons sublist'_sublist)
-done
-
-
-subsection {* Showing equivalence to use of drop and take for definition *}
-
-lemma "sublist' n m xs = take (m - n) (drop n xs)"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (case_tac n)
-apply simp
-apply simp
-done
-
-subsection {* General lemma about sublist *}
-
-lemma sublist'_Nil[simp]: "sublist' i j [] = []"
-by simp
-
-lemma sublist'_Cons[simp]: "sublist' i (Suc j) (x#xs) = (case i of 0 \<Rightarrow> (x # sublist' 0 j xs) | Suc i' \<Rightarrow>  sublist' i' j xs)"
-by (cases i) auto
-
-lemma sublist'_Cons2[simp]: "sublist' i j (x#xs) = (if (j = 0) then [] else ((if (i = 0) then [x] else []) @ sublist' (i - 1) (j - 1) xs))"
-apply (cases j)
-apply auto
-apply (cases i)
-apply auto
-done
-
-lemma sublist_n_0: "sublist' n 0 xs = []"
-by (induct xs, auto)
-
-lemma sublist'_Nil': "n \<ge> m \<Longrightarrow> sublist' n m xs = []"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (case_tac n)
-apply simp
-apply simp
-done
-
-lemma sublist'_Nil2: "n \<ge> length xs \<Longrightarrow> sublist' n m xs = []"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (case_tac n)
-apply simp
-apply simp
-done
-
-lemma sublist'_Nil3: "(sublist' n m xs = []) = ((n \<ge> m) \<or> (n \<ge> length xs))"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (case_tac n)
-apply simp
-apply simp
-done
-
-lemma sublist'_notNil: "\<lbrakk> n < length xs; n < m \<rbrakk> \<Longrightarrow> sublist' n m xs \<noteq> []"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (case_tac n)
-apply simp
-apply simp
-done
-
-lemma sublist'_single: "n < length xs \<Longrightarrow> sublist' n (Suc n) xs = [xs ! n]"
-apply (induct xs arbitrary: n)
-apply simp
-apply simp
-apply (case_tac n)
-apply (simp add: sublist_n_0)
-apply simp
-done
-
-lemma sublist'_update1: "i \<ge> m \<Longrightarrow> sublist' n m (xs[i:=v]) = sublist' n m xs"
-apply (induct xs arbitrary: n m i)
-apply simp
-apply simp
-apply (case_tac i)
-apply simp
-apply simp
-done
-
-lemma sublist'_update2: "i < n \<Longrightarrow> sublist' n m (xs[i:=v]) = sublist' n m xs"
-apply (induct xs arbitrary: n m i)
-apply simp
-apply simp
-apply (case_tac i)
-apply simp
-apply simp
-done
-
-lemma sublist'_update3: "\<lbrakk>n \<le> i; i < m\<rbrakk> \<Longrightarrow> sublist' n m (xs[i := v]) = (sublist' n m xs)[i - n := v]"
-proof (induct xs arbitrary: n m i)
-  case Nil thus ?case by auto
-next
-  case (Cons x xs)
-  thus ?case
-    apply -
-    apply auto
-    apply (cases i)
-    apply auto
-    apply (cases i)
-    apply auto
-    done
-qed
-
-lemma "\<lbrakk> sublist' i j xs = sublist' i j ys; n \<ge> i; m \<le> j \<rbrakk> \<Longrightarrow> sublist' n m xs = sublist' n m ys"
-proof (induct xs arbitrary: i j ys n m)
-  case Nil
-  thus ?case
-    apply -
-    apply (rule sym, drule sym)
-    apply (simp add: sublist'_Nil)
-    apply (simp add: sublist'_Nil3)
-    apply arith
-    done
-next
-  case (Cons x xs i j ys n m)
-  note c = this
-  thus ?case
-  proof (cases m)
-    case 0 thus ?thesis by (simp add: sublist_n_0)
-  next
-    case (Suc m')
-    note a = this
-    thus ?thesis
-    proof (cases n)
-      case 0 note b = this
-      show ?thesis
-      proof (cases ys)
-	case Nil  with a b Cons.prems show ?thesis by (simp add: sublist'_Nil3)
-      next
-	case (Cons y ys)
-	show ?thesis
-	proof (cases j)
-	  case 0 with a b Cons.prems show ?thesis by simp
-	next
-	  case (Suc j') with a b Cons.prems Cons show ?thesis 
-	    apply -
-	    apply (simp, rule Cons.hyps [of "0" "j'" "ys" "0" "m'"], auto)
-	    done
-	qed
-      qed
-    next
-      case (Suc n')
-      show ?thesis
-      proof (cases ys)
-	case Nil with Suc a Cons.prems show ?thesis by (auto simp add: sublist'_Nil3)
-      next
-	case (Cons y ys) with Suc a Cons.prems show ?thesis
-	  apply -
-	  apply simp
-	  apply (cases j)
-	  apply simp
-	  apply (cases i)
-	  apply simp
-	  apply (rule_tac j="nat" in Cons.hyps [of "0" _ "ys" "n'" "m'"])
-	  apply simp
-	  apply simp
-	  apply simp
-	  apply simp
-	  apply (rule_tac i="nata" and j="nat" in Cons.hyps [of _ _ "ys" "n'" "m'"])
-	  apply simp
-	  apply simp
-	  apply simp
-	  done
-      qed
-    qed
-  qed
-qed
-
-lemma length_sublist': "j \<le> length xs \<Longrightarrow> length (sublist' i j xs) = j - i"
-by (induct xs arbitrary: i j, auto)
-
-lemma sublist'_front: "\<lbrakk> i < j; i < length xs \<rbrakk> \<Longrightarrow> sublist' i j xs = xs ! i # sublist' (Suc i) j xs"
-apply (induct xs arbitrary: a i j)
-apply simp
-apply (case_tac j)
-apply simp
-apply (case_tac i)
-apply simp
-apply simp
-done
-
-lemma sublist'_back: "\<lbrakk> i < j; j \<le> length xs \<rbrakk> \<Longrightarrow> sublist' i j xs = sublist' i (j - 1) xs @ [xs ! (j - 1)]"
-apply (induct xs arbitrary: a i j)
-apply simp
-apply simp
-apply (case_tac j)
-apply simp
-apply auto
-apply (case_tac nat)
-apply auto
-done
-
-(* suffices that j \<le> length xs and length ys *) 
-lemma sublist'_eq_samelength_iff: "length xs = length ys \<Longrightarrow> (sublist' i j xs  = sublist' i j ys) = (\<forall>i'. i \<le> i' \<and> i' < j \<longrightarrow> xs ! i' = ys ! i')"
-proof (induct xs arbitrary: ys i j)
-  case Nil thus ?case by simp
-next
-  case (Cons x xs)
-  thus ?case
-    apply -
-    apply (cases ys)
-    apply simp
-    apply simp
-    apply auto
-    apply (case_tac i', auto)
-    apply (erule_tac x="Suc i'" in allE, auto)
-    apply (erule_tac x="i' - 1" in allE, auto)
-    apply (case_tac i', auto)
-    apply (erule_tac x="Suc i'" in allE, auto)
-    done
-qed
-
-lemma sublist'_all[simp]: "sublist' 0 (length xs) xs = xs"
-by (induct xs, auto)
-
-lemma sublist'_sublist': "sublist' n m (sublist' i j xs) = sublist' (i + n) (min (i + m) j) xs" 
-by (induct xs arbitrary: i j n m) (auto simp add: min_diff)
-
-lemma sublist'_append: "\<lbrakk> i \<le> j; j \<le> k \<rbrakk> \<Longrightarrow>(sublist' i j xs) @ (sublist' j k xs) = sublist' i k xs"
-by (induct xs arbitrary: i j k) auto
-
-lemma nth_sublist': "\<lbrakk> k < j - i; j \<le> length xs \<rbrakk> \<Longrightarrow> (sublist' i j xs) ! k = xs ! (i + k)"
-apply (induct xs arbitrary: i j k)
-apply auto
-apply (case_tac k)
-apply auto
-apply (case_tac i)
-apply auto
-done
-
-lemma set_sublist': "set (sublist' i j xs) = {x. \<exists>k. i \<le> k \<and> k < j \<and> k < List.length xs \<and> x = xs ! k}"
-apply (simp add: sublist'_sublist)
-apply (simp add: set_sublist)
-apply auto
-done
-
-lemma all_in_set_sublist'_conv: "(\<forall>j. j \<in> set (sublist' l r xs) \<longrightarrow> P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < List.length xs \<longrightarrow> P (xs ! k))"
-unfolding set_sublist' by blast
-
-lemma ball_in_set_sublist'_conv: "(\<forall>j \<in> set (sublist' l r xs). P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < List.length xs \<longrightarrow> P (xs ! k))"
-unfolding set_sublist' by blast
-
-
-lemma multiset_of_sublist:
-assumes l_r: "l \<le> r \<and> r \<le> List.length xs"
-assumes left: "\<forall> i. i < l \<longrightarrow> (xs::'a list) ! i = ys ! i"
-assumes right: "\<forall> i. i \<ge> r \<longrightarrow> (xs::'a list) ! i = ys ! i"
-assumes multiset: "multiset_of xs = multiset_of ys"
-  shows "multiset_of (sublist' l r xs) = multiset_of (sublist' l r ys)"
-proof -
-  from l_r have xs_def: "xs = (sublist' 0 l xs) @ (sublist' l r xs) @ (sublist' r (List.length xs) xs)" (is "_ = ?xs_long") 
-    by (simp add: sublist'_append)
-  from multiset have length_eq: "List.length xs = List.length ys" by (rule multiset_of_eq_length)
-  with l_r have ys_def: "ys = (sublist' 0 l ys) @ (sublist' l r ys) @ (sublist' r (List.length ys) ys)" (is "_ = ?ys_long") 
-    by (simp add: sublist'_append)
-  from xs_def ys_def multiset have "multiset_of ?xs_long = multiset_of ?ys_long" by simp
-  moreover
-  from left l_r length_eq have "sublist' 0 l xs = sublist' 0 l ys"
-    by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI)
-  moreover
-  from right l_r length_eq have "sublist' r (List.length xs) xs = sublist' r (List.length ys) ys"
-    by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI)
-  moreover
-  ultimately show ?thesis by (simp add: multiset_of_append)
-qed
-
-
-end
--- a/src/HOLCF/Algebraic.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/Algebraic.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -215,7 +215,7 @@
 lemma finite_deflation_Rep_fin_defl: "finite_deflation (Rep_fin_defl d)"
 using Rep_fin_defl by simp
 
-interpretation Rep_fin_defl!: finite_deflation "Rep_fin_defl d"
+interpretation Rep_fin_defl: finite_deflation "Rep_fin_defl d"
 by (rule finite_deflation_Rep_fin_defl)
 
 lemma fin_defl_lessI:
@@ -320,7 +320,7 @@
 apply (rule Rep_fin_defl.compact)
 done
 
-interpretation fin_defl!: basis_take sq_le fd_take
+interpretation fin_defl: basis_take sq_le fd_take
 apply default
 apply (rule fd_take_less)
 apply (rule fd_take_idem)
@@ -370,7 +370,7 @@
 unfolding alg_defl_principal_def
 by (simp add: Abs_alg_defl_inverse sq_le.ideal_principal)
 
-interpretation alg_defl!:
+interpretation alg_defl:
   ideal_completion sq_le fd_take alg_defl_principal Rep_alg_defl
 apply default
 apply (rule ideal_Rep_alg_defl)
@@ -461,7 +461,7 @@
 apply (rule finite_deflation_Rep_fin_defl)
 done
 
-interpretation cast!: deflation "cast\<cdot>d"
+interpretation cast: deflation "cast\<cdot>d"
 by (rule deflation_cast)
 
 lemma "cast\<cdot>(\<Squnion>i. alg_defl_principal (Abs_fin_defl (approx i)))\<cdot>x = x"
--- a/src/HOLCF/Bifinite.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/Bifinite.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -37,7 +37,7 @@
     by (rule finite_fixes_approx)
 qed
 
-interpretation approx!: finite_deflation "approx i"
+interpretation approx: finite_deflation "approx i"
 by (rule finite_deflation_approx)
 
 lemma (in deflation) deflation: "deflation d" ..
--- a/src/HOLCF/CompactBasis.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/CompactBasis.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -46,7 +46,7 @@
 
 lemmas approx_Rep_compact_basis = Rep_compact_take [symmetric]
 
-interpretation compact_basis!:
+interpretation compact_basis:
   basis_take sq_le compact_take
 proof
   fix n :: nat and a :: "'a compact_basis"
@@ -92,7 +92,7 @@
   approximants :: "'a \<Rightarrow> 'a compact_basis set" where
   "approximants = (\<lambda>x. {a. Rep_compact_basis a \<sqsubseteq> x})"
 
-interpretation compact_basis!:
+interpretation compact_basis:
   ideal_completion sq_le compact_take Rep_compact_basis approximants
 proof
   fix w :: 'a
--- a/src/HOLCF/Completion.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/Completion.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -156,7 +156,7 @@
 
 end
 
-interpretation sq_le!: preorder "sq_le :: 'a::po \<Rightarrow> 'a \<Rightarrow> bool"
+interpretation sq_le: preorder "sq_le :: 'a::po \<Rightarrow> 'a \<Rightarrow> bool"
 apply unfold_locales
 apply (rule refl_less)
 apply (erule (1) trans_less)
--- a/src/HOLCF/ConvexPD.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/ConvexPD.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -20,7 +20,7 @@
 lemma convex_le_trans: "\<lbrakk>t \<le>\<natural> u; u \<le>\<natural> v\<rbrakk> \<Longrightarrow> t \<le>\<natural> v"
 unfolding convex_le_def by (fast intro: upper_le_trans lower_le_trans)
 
-interpretation convex_le!: preorder convex_le
+interpretation convex_le: preorder convex_le
 by (rule preorder.intro, rule convex_le_refl, rule convex_le_trans)
 
 lemma upper_le_minimal [simp]: "PDUnit compact_bot \<le>\<natural> t"
@@ -178,7 +178,7 @@
 unfolding convex_principal_def
 by (simp add: Abs_convex_pd_inverse convex_le.ideal_principal)
 
-interpretation convex_pd!:
+interpretation convex_pd:
   ideal_completion convex_le pd_take convex_principal Rep_convex_pd
 apply unfold_locales
 apply (rule pd_take_convex_le)
--- a/src/HOLCF/FOCUS/Buffer.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/FOCUS/Buffer.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -98,8 +98,11 @@
 by (erule subst, rule refl)
 
 ML {*
-fun B_prover s tac eqs = prove_goal (the_context ()) s (fn prems => [cut_facts_tac prems 1,
-        tac 1, auto_tac (claset(), simpset() addsimps eqs)]);
+fun B_prover s tac eqs =
+  let val thy = the_context () in
+    prove_goal thy s (fn prems => [cut_facts_tac prems 1,
+        tac 1, auto_tac (claset_of thy, simpset_of thy addsimps eqs)])
+  end;
 
 fun prove_forw  s thm     = B_prover s (dtac (thm RS iffD1)) [];
 fun prove_backw s thm eqs = B_prover s (rtac (thm RS iffD2)) eqs;
--- a/src/HOLCF/HOLCF.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/HOLCF.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -24,7 +24,8 @@
 declaration {* fn _ =>
   Simplifier.map_ss (fn simpset => simpset addSolver
     (mk_solver' "adm_tac" (fn ss =>
-      Adm.adm_tac (cut_facts_tac (Simplifier.prems_of_ss ss) THEN' cont_tacRs ss))));
+      Adm.adm_tac (Simplifier.the_context ss)
+        (cut_facts_tac (Simplifier.prems_of_ss ss) THEN' cont_tacRs ss))));
 *}
 
 end
--- a/src/HOLCF/IOA/Modelcheck/Cockpit.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/IOA/Modelcheck/Cockpit.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -102,18 +102,18 @@
 
 (* to prove, that info is always set at the recent alarm *)
 lemma cockpit_implements_Info_while_Al: "cockpit =<| Info_while_Al"
-apply (tactic {* is_sim_tac (thms "aut_simps") 1 *})
+apply (tactic {* is_sim_tac @{simpset} (thms "aut_simps") 1 *})
 done
 
 (* to prove that before any alarm arrives (and after each acknowledgment),
    info remains at None *)
 lemma cockpit_implements_Info_before_Al: "cockpit =<| Info_before_Al"
-apply (tactic {* is_sim_tac (thms "aut_simps") 1 *})
+apply (tactic {* is_sim_tac @{simpset} (thms "aut_simps") 1 *})
 done
 
 (* to prove that before any alarm would be acknowledged, it must be arrived *)
 lemma cockpit_implements_Al_before_Ack: "cockpit_hide =<| Al_before_Ack"
-apply (tactic {* is_sim_tac (thms "aut_simps") 1 *})
+apply (tactic {* is_sim_tac @{simpset} (thms "aut_simps") 1 *})
 apply auto
 done
 
--- a/src/HOLCF/IOA/Modelcheck/MuIOA.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/IOA/Modelcheck/MuIOA.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -278,14 +278,14 @@
 by (REPEAT (rtac impI 1));
 by (REPEAT (dtac eq_reflection 1));
 (* Bis hierher wird im Kapitel 4 erl"autert, ab hier im Kapitel 5 *)
-by (full_simp_tac ((simpset() delcongs (if_weak_cong :: weak_case_congs)
+by (full_simp_tac ((simpset_of sign delcongs (if_weak_cong :: weak_case_congs)
                               delsimps [not_iff,split_part])
 			      addsimps (thl @ comp_simps @ restrict_simps @ hide_simps @
                                         rename_simps @ ioa_simps @ asig_simps)) 1);
 by (full_simp_tac
   (Mucke_ss delcongs (if_weak_cong :: weak_case_congs) delsimps [not_iff,split_part]) 1);
 by (REPEAT (if_full_simp_tac
-  (simpset() delcongs (if_weak_cong :: weak_case_congs) delsimps [not_iff,split_part]) 1));
+  (simpset_of sign delcongs (if_weak_cong :: weak_case_congs) delsimps [not_iff,split_part]) 1));
 by (call_mucke_tac 1);
 (* Bis hierher wird im Kapitel 5 erl"autert, ab hier wieder im Kapitel 4 *)
 by (atac 1);
--- a/src/HOLCF/IOA/Modelcheck/MuIOAOracle.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/IOA/Modelcheck/MuIOAOracle.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -18,18 +18,18 @@
 
 (* is_sim_tac makes additionally to call_sim_tac some simplifications,
 	which are suitable for implementation realtion formulas *)
-fun is_sim_tac thm_list = SUBGOAL (fn (goal, i) =>
+fun is_sim_tac ss thm_list = SUBGOAL (fn (goal, i) =>
   EVERY [REPEAT (etac thin_rl i),
-	 simp_tac (simpset() addsimps [ioa_implements_def]) i,
+	 simp_tac (ss addsimps [ioa_implements_def]) i,
          rtac conjI i,
          rtac conjI (i+1),
 	 TRY(call_sim_tac thm_list (i+2)),
 	 TRY(atac (i+2)), 
          REPEAT(rtac refl (i+2)),
-	 simp_tac (simpset() addsimps (thm_list @
+	 simp_tac (ss addsimps (thm_list @
 				       comp_simps @ restrict_simps @ hide_simps @
 				       rename_simps @  ioa_simps @ asig_simps)) (i+1),
-	 simp_tac (simpset() addsimps (thm_list @
+	 simp_tac (ss addsimps (thm_list @
 				       comp_simps @ restrict_simps @ hide_simps @
 				       rename_simps @ ioa_simps @ asig_simps)) (i)]);
 
--- a/src/HOLCF/IOA/Modelcheck/Ring3.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/IOA/Modelcheck/Ring3.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -114,7 +114,7 @@
 (* property to prove: at most one (of 3) members of the ring will
 	declare itself to leader *)
 lemma at_most_one_leader: "ring =<| one_leader"
-apply (tactic {* is_sim_tac (thms "aut_simps") 1 *})
+apply (tactic {* is_sim_tac @{simpset} (thms "aut_simps") 1 *})
 apply auto
 done
 
--- a/src/HOLCF/IOA/ex/TrivEx.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/IOA/ex/TrivEx.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOLCF/IOA/TrivEx.thy
-    ID:         $Id$
     Author:     Olaf Müller
 *)
 
@@ -66,7 +65,7 @@
 apply (rule AbsRuleT1)
 apply (rule h_abs_is_abstraction)
 apply (rule MC_result)
-apply (tactic "abstraction_tac 1")
+apply abstraction
 apply (simp add: h_abs_def)
 done
 
--- a/src/HOLCF/IOA/ex/TrivEx2.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/IOA/ex/TrivEx2.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOLCF/IOA/TrivEx.thy
-    ID:         $Id$
     Author:     Olaf Müller
 *)
 
@@ -85,7 +84,7 @@
 txt {* is_abstraction *}
 apply (rule h_abs_is_abstraction)
 txt {* temp_weakening *}
-apply (tactic "abstraction_tac 1")
+apply abstraction
 apply (erule Enabled_implication)
 done
 
@@ -96,7 +95,7 @@
 apply (rule AbsRuleT2)
 apply (rule h_abs_is_liveabstraction)
 apply (rule MC_result)
-apply (tactic "abstraction_tac 1")
+apply abstraction
 apply (simp add: h_abs_def)
 done
 
--- a/src/HOLCF/IOA/meta_theory/Abstraction.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/IOA/meta_theory/Abstraction.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOLCF/IOA/meta_theory/Abstraction.thy
-    ID:         $Id$
     Author:     Olaf Müller
 *)
 
@@ -605,12 +604,15 @@
   strength_Diamond strength_Leadsto weak_WF weak_SF
 
 ML {*
-fun abstraction_tac i =
-    SELECT_GOAL (CLASIMPSET (fn (cs, ss) =>
-      auto_tac (cs addSIs @{thms weak_strength_lemmas},
-        ss addsimps [@{thm state_strengthening_def}, @{thm state_weakening_def}]))) i
+fun abstraction_tac ctxt =
+  let val (cs, ss) = local_clasimpset_of ctxt in
+    SELECT_GOAL (auto_tac (cs addSIs @{thms weak_strength_lemmas},
+        ss addsimps [@{thm state_strengthening_def}, @{thm state_weakening_def}]))
+  end
 *}
 
+method_setup abstraction = {* Scan.succeed (SIMPLE_METHOD' o abstraction_tac) *} ""
+
 use "ioa_package.ML"
 
 end
--- a/src/HOLCF/IOA/meta_theory/Sequence.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/IOA/meta_theory/Sequence.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOLCF/IOA/meta_theory/Sequence.thy
-    ID:         $Id$
     Author:     Olaf Müller
 
 Sequences over flat domains with lifted elements.
@@ -340,7 +339,7 @@
 lemma Seq_induct:
 "!! P. [| adm P; P UU; P nil; !! a s. P s ==> P (a>>s)|] ==> P x"
 apply (erule (2) seq.ind)
-apply (tactic {* def_tac 1 *})
+apply defined
 apply (simp add: Consq_def)
 done
 
@@ -348,14 +347,14 @@
 "!! P.[|P UU;P nil; !! a s. P s ==> P(a>>s) |]
                 ==> seq_finite x --> P x"
 apply (erule (1) seq_finite_ind)
-apply (tactic {* def_tac 1 *})
+apply defined
 apply (simp add: Consq_def)
 done
 
 lemma Seq_Finite_ind:
 "!! P.[| Finite x; P nil; !! a s. [| Finite s; P s|] ==> P (a>>s) |] ==> P x"
 apply (erule (1) Finite.induct)
-apply (tactic {* def_tac 1 *})
+apply defined
 apply (simp add: Consq_def)
 done
 
--- a/src/HOLCF/IOA/meta_theory/Traces.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/IOA/meta_theory/Traces.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOLCF/IOA/meta_theory/Traces.thy
-    ID:         $Id$
     Author:     Olaf Müller
 *)
 
@@ -327,7 +326,7 @@
 apply (simp (no_asm_simp))
 apply (drule Finite_Last1 [THEN mp])
 apply assumption
-apply (tactic "def_tac 1")
+apply defined
 done
 
 declare laststate_UU [simp] laststate_nil [simp] laststate_cons [simp]
--- a/src/HOLCF/Lift.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/Lift.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -56,15 +56,13 @@
   by (cases x) simp_all
 
 text {*
-  For @{term "x ~= UU"} in assumptions @{text def_tac} replaces @{text
+  For @{term "x ~= UU"} in assumptions @{text defined} replaces @{text
   x} by @{text "Def a"} in conclusion. *}
 
-ML {*
-  local val lift_definedE = thm "lift_definedE"
-  in val def_tac = SIMPSET' (fn ss =>
-    etac lift_definedE THEN' asm_simp_tac ss)
-  end;
-*}
+method_setup defined = {*
+  Scan.succeed (fn ctxt => SIMPLE_METHOD'
+    (etac @{thm lift_definedE} THEN' asm_simp_tac (local_simpset_of ctxt)))
+*} ""
 
 lemma DefE: "Def x = \<bottom> \<Longrightarrow> R"
   by simp
--- a/src/HOLCF/LowerPD.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/LowerPD.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -26,7 +26,7 @@
 apply (erule (1) trans_less)
 done
 
-interpretation lower_le!: preorder lower_le
+interpretation lower_le: preorder lower_le
 by (rule preorder.intro, rule lower_le_refl, rule lower_le_trans)
 
 lemma lower_le_minimal [simp]: "PDUnit compact_bot \<le>\<flat> t"
@@ -133,7 +133,7 @@
 unfolding lower_principal_def
 by (simp add: Abs_lower_pd_inverse lower_le.ideal_principal)
 
-interpretation lower_pd!:
+interpretation lower_pd:
   ideal_completion lower_le pd_take lower_principal Rep_lower_pd
 apply unfold_locales
 apply (rule pd_take_lower_le)
--- a/src/HOLCF/Tools/adm_tac.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/Tools/adm_tac.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -15,7 +15,7 @@
 
 signature ADM =
 sig
-  val adm_tac: (int -> tactic) -> int -> tactic
+  val adm_tac: Proof.context -> (int -> tactic) -> int -> tactic
 end;
 
 structure Adm: ADM =
@@ -85,14 +85,14 @@
 (*** try to prove that terms in list are continuous
      if successful, add continuity theorem to list l ***)
 
-fun prove_cont tac thy s T prems params (ts as ((t, _)::_)) l =
+fun prove_cont ctxt tac s T prems params (ts as ((t, _)::_)) l =  (* FIXME proper context *)
   let val parTs = map snd (rev params);
        val contT = (T --> (fastype_of1 (T::parTs, t))) --> HOLogic.boolT;
        fun mk_all [] t = t
          | mk_all ((a,T)::Ts) t = Term.all T $ (Abs (a, T, mk_all Ts t));
        val t = HOLogic.mk_Trueprop (Const (@{const_name cont}, contT) $ Abs (s, T, t));
        val t' = mk_all params (Logic.list_implies (prems, t));
-       val thm = Goal.prove (ProofContext.init thy) [] [] t' (K (tac 1));
+       val thm = Goal.prove ctxt [] [] t' (K (tac 1));
   in (ts, thm)::l end
   handle ERROR _ => l;
 
@@ -128,18 +128,18 @@
 fun try_dest_adm (Const _ $ (Const (@{const_name adm}, _) $ Abs abs)) = SOME abs
   | try_dest_adm _ = NONE;
 
-fun adm_tac tac i state = (i, state) |-> SUBGOAL (fn (goali, _) =>
+fun adm_tac ctxt tac i state = (i, state) |-> SUBGOAL (fn (goali, _) =>
   (case try_dest_adm (Logic.strip_assums_concl goali) of
     NONE => no_tac
   | SOME (s, T, t) =>
       let
-        val thy = Thm.theory_of_thm state;
+        val thy = ProofContext.theory_of ctxt;
         val prems = Logic.strip_assums_hyp goali;
         val params = Logic.strip_params goali;
         val ts = find_subterms t 0 [];
         val ts' = filter eq_terms ts;
         val ts'' = filter (is_chfin thy T params) ts';
-        val thms = fold (prove_cont tac thy s T prems params) ts'' [];
+        val thms = fold (prove_cont ctxt tac s T prems params) ts'' [];
       in
         (case thms of
           ((ts as ((t', _)::_), cont_thm) :: _) =>
--- a/src/HOLCF/Tools/domain/domain_library.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/Tools/domain/domain_library.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -125,37 +125,37 @@
 infix 0 ==;     fun S ==  T = %%:"==" $ S $ T;
 infix 1 ===;    fun S === T = %%:"op =" $ S $ T;
 infix 1 ~=;     fun S ~=  T = HOLogic.mk_not (S === T);
-infix 1 <<;     fun S <<  T = %%:@{const_name Porder.sq_le} $ S $ T;
+infix 1 <<;     fun S <<  T = %%: @{const_name Porder.sq_le} $ S $ T;
 infix 1 ~<<;    fun S ~<< T = HOLogic.mk_not (S << T);
 
-infix 9 `  ; fun f ` x = %%:@{const_name Rep_CFun} $ f $ x;
+infix 9 `  ; fun f ` x = %%: @{const_name Rep_CFun} $ f $ x;
 infix 9 `% ; fun f`% s = f` %: s;
 infix 9 `%%; fun f`%%s = f` %%:s;
 
-fun mk_adm t = %%:@{const_name adm} $ t;
-fun mk_compact t = %%:@{const_name compact} $ t;
-val ID = %%:@{const_name ID};
-fun mk_strictify t = %%:@{const_name strictify}`t;
-fun mk_cfst t = %%:@{const_name cfst}`t;
-fun mk_csnd t = %%:@{const_name csnd}`t;
+fun mk_adm t = %%: @{const_name adm} $ t;
+fun mk_compact t = %%: @{const_name compact} $ t;
+val ID = %%: @{const_name ID};
+fun mk_strictify t = %%: @{const_name strictify}`t;
+fun mk_cfst t = %%: @{const_name cfst}`t;
+fun mk_csnd t = %%: @{const_name csnd}`t;
 (*val csplitN    = "Cprod.csplit";*)
 (*val sfstN      = "Sprod.sfst";*)
 (*val ssndN      = "Sprod.ssnd";*)
-fun mk_ssplit t = %%:@{const_name ssplit}`t;
-fun mk_sinl t = %%:@{const_name sinl}`t;
-fun mk_sinr t = %%:@{const_name sinr}`t;
-fun mk_sscase (x, y) = %%:@{const_name sscase}`x`y;
-fun mk_up t = %%:@{const_name up}`t;
-fun mk_fup (t,u) = %%:@{const_name fup} ` t ` u;
+fun mk_ssplit t = %%: @{const_name ssplit}`t;
+fun mk_sinl t = %%: @{const_name sinl}`t;
+fun mk_sinr t = %%: @{const_name sinr}`t;
+fun mk_sscase (x, y) = %%: @{const_name sscase}`x`y;
+fun mk_up t = %%: @{const_name up}`t;
+fun mk_fup (t,u) = %%: @{const_name fup} ` t ` u;
 val ONE = @{term ONE};
 val TT = @{term TT};
 val FF = @{term FF};
-fun mk_iterate (n,f,z) = %%:@{const_name iterate} $ n ` f ` z;
-fun mk_fix t = %%:@{const_name fix}`t;
-fun mk_return t = %%:@{const_name Fixrec.return}`t;
-val mk_fail = %%:@{const_name Fixrec.fail};
+fun mk_iterate (n,f,z) = %%: @{const_name iterate} $ n ` f ` z;
+fun mk_fix t = %%: @{const_name fix}`t;
+fun mk_return t = %%: @{const_name Fixrec.return}`t;
+val mk_fail = %%: @{const_name Fixrec.fail};
 
-fun mk_branch t = %%:@{const_name Fixrec.branch} $ t;
+fun mk_branch t = %%: @{const_name Fixrec.branch} $ t;
 
 val pcpoS = @{sort pcpo};
 
@@ -171,14 +171,14 @@
 fun cproj x      = prj (fn S => K(mk_cfst S)) (fn S => K(mk_csnd S)) x;
 fun lift tfn = Library.foldr (fn (x,t)=> (mk_trp(tfn x) ===> t));
 
-fun /\ v T = %%:@{const_name Abs_CFun} $ mk_lam(v,T);
+fun /\ v T = %%: @{const_name Abs_CFun} $ mk_lam(v,T);
 fun /\# (arg,T) = /\ (vname arg) T;
-infixr 9 oo; fun S oo T = %%:@{const_name cfcomp}`S`T;
-val UU = %%:@{const_name UU};
+infixr 9 oo; fun S oo T = %%: @{const_name cfcomp}`S`T;
+val UU = %%: @{const_name UU};
 fun strict f = f`UU === UU;
 fun defined t = t ~= UU;
-fun cpair (t,u) = %%:@{const_name cpair}`t`u;
-fun spair (t,u) = %%:@{const_name spair}`t`u;
+fun cpair (t,u) = %%: @{const_name cpair}`t`u;
+fun spair (t,u) = %%: @{const_name spair}`t`u;
 fun mk_ctuple [] = HOLogic.unit (* used in match_defs *)
 |   mk_ctuple ts = foldr1 cpair ts;
 fun mk_stuple [] = ONE
@@ -186,7 +186,7 @@
 fun mk_ctupleT [] = HOLogic.unitT   (* used in match_defs *)
 |   mk_ctupleT Ts = foldr1 HOLogic.mk_prodT Ts;
 fun mk_maybeT T = Type ("Fixrec.maybe",[T]);
-fun cpair_pat (p1,p2) = %%:@{const_name cpair_pat} $ p1 $ p2;
+fun cpair_pat (p1,p2) = %%: @{const_name cpair_pat} $ p1 $ p2;
 val mk_ctuple_pat = foldr1 cpair_pat;
 fun lift_defined f = lift (fn x => defined (f x));
 fun bound_arg vns v = Bound(length vns -find_index_eq v vns -1);
--- a/src/HOLCF/Universal.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/Universal.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -230,13 +230,13 @@
 apply (simp add: ubasis_take_same)
 done
 
-interpretation udom!: preorder ubasis_le
+interpretation udom: preorder ubasis_le
 apply default
 apply (rule ubasis_le_refl)
 apply (erule (1) ubasis_le_trans)
 done
 
-interpretation udom!: basis_take ubasis_le ubasis_take
+interpretation udom: basis_take ubasis_le ubasis_take
 apply default
 apply (rule ubasis_take_less)
 apply (rule ubasis_take_idem)
@@ -285,7 +285,7 @@
 unfolding udom_principal_def
 by (simp add: Abs_udom_inverse udom.ideal_principal)
 
-interpretation udom!:
+interpretation udom:
   ideal_completion ubasis_le ubasis_take udom_principal Rep_udom
 apply unfold_locales
 apply (rule ideal_Rep_udom)
--- a/src/HOLCF/UpperPD.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/HOLCF/UpperPD.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -26,7 +26,7 @@
 apply (erule (1) trans_less)
 done
 
-interpretation upper_le!: preorder upper_le
+interpretation upper_le: preorder upper_le
 by (rule preorder.intro, rule upper_le_refl, rule upper_le_trans)
 
 lemma upper_le_minimal [simp]: "PDUnit compact_bot \<le>\<sharp> t"
@@ -131,7 +131,7 @@
 unfolding upper_principal_def
 by (simp add: Abs_upper_pd_inverse upper_le.ideal_principal)
 
-interpretation upper_pd!:
+interpretation upper_pd:
   ideal_completion upper_le pd_take upper_principal Rep_upper_pd
 apply unfold_locales
 apply (rule pd_take_upper_le)
--- a/src/Provers/Arith/assoc_fold.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Provers/Arith/assoc_fold.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -12,6 +12,7 @@
 sig
   val assoc_ss: simpset
   val eq_reflection: thm
+  val is_numeral: term -> bool
 end;
 
 signature ASSOC_FOLD =
@@ -29,10 +30,9 @@
 
 (*Separate the literals from the other terms being combined*)
 fun sift_terms plus (t, (lits,others)) =
+  if Data.is_numeral t then (t::lits, others) (*new literal*) else
   (case t of
-    Const (@{const_name Int.number_of}, _) $ _ =>       (* FIXME logic dependent *)
-      (t::lits, others)         (*new literal*)
-  | (f as Const _) $ x $ y =>
+    (f as Const _) $ x $ y =>
       if f = plus
       then sift_terms plus (x, sift_terms plus (y, (lits,others)))
       else (lits, t::others)    (*arbitrary summand*)
--- a/src/Provers/Arith/extract_common_term.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Provers/Arith/extract_common_term.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -25,7 +25,8 @@
   (*proof tools*)
   val prove_conv: tactic list -> Proof.context -> thm list -> term * term -> thm option
   val norm_tac: simpset -> tactic                (*proves the result*)
-  val simplify_meta_eq: simpset -> thm -> thm    (*simplifies the result*)
+  val simplify_meta_eq: simpset -> thm -> thm -> thm (*simplifies the result*)
+  val simp_conv: simpset -> term -> thm option  (*proves simp thm*)
 end;
 
 
@@ -58,6 +59,9 @@
     and terms2 = Data.dest_sum t2
 
     val u = find_common (terms1,terms2)
+    val simp_th =
+          case Data.simp_conv ss u of NONE => raise TERM("no simp", [])
+          | SOME th => th
     val terms1' = Data.find_first u terms1
     and terms2' = Data.find_first u terms2
     and T = Term.fastype_of u
@@ -66,7 +70,7 @@
       Data.prove_conv [Data.norm_tac ss] ctxt prems
         (t', Data.mk_bal (Data.mk_sum T (u::terms1'), Data.mk_sum T (u::terms2')))
   in
-    Option.map (export o Data.simplify_meta_eq ss) reshape
+    Option.map (export o Data.simplify_meta_eq ss simp_th) reshape
   end
   (* FIXME avoid handling of generic exceptions *)
   handle TERM _ => NONE
--- a/src/Provers/Arith/fast_lin_arith.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Provers/Arith/fast_lin_arith.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -466,7 +466,7 @@
                      NONE => ( the (try_add ([thm2] RL inj_thms) thm1)
                                handle Option =>
                                (trace_thm "" thm1; trace_thm "" thm2;
-                                sys_error "Lin.arith. failed to add thms")
+                                sys_error "Linear arithmetic: failed to add thms")
                              )
                    | SOME thm => thm)
         | SOME thm => thm;
@@ -588,8 +588,8 @@
           handle NoEx => NONE
       in
         case ex of
-          SOME s => (warning "arith failed - see trace for a counterexample"; tracing s)
-        | NONE => warning "arith failed"
+          SOME s => (warning "Linear arithmetic failed - see trace for a counterexample."; tracing s)
+        | NONE => warning "Linear arithmetic failed"
       end;
 
 (* ------------------------------------------------------------------------- *)
--- a/src/Provers/blast.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Provers/blast.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -48,7 +48,6 @@
   val contr_tac         : int -> tactic
   val dup_intr          : thm -> thm
   val hyp_subst_tac     : bool -> int -> tactic
-  val claset            : unit -> claset
   val rep_cs    : (* dependent on classical.ML *)
       claset -> {safeIs: thm list, safeEs: thm list,
                  hazIs: thm list, hazEs: thm list,
@@ -77,7 +76,6 @@
   val depth_tac         : claset -> int -> int -> tactic
   val depth_limit       : int Config.T
   val blast_tac         : claset -> int -> tactic
-  val Blast_tac         : int -> tactic
   val setup             : theory -> theory
   (*debugging tools*)
   val stats             : bool ref
@@ -89,7 +87,7 @@
   val instVars          : term -> (unit -> unit)
   val toTerm            : int -> term -> Term.term
   val readGoal          : theory -> string -> term
-  val tryInThy          : theory -> int -> string ->
+  val tryInThy          : theory -> claset -> int -> string ->
                   (int->tactic) list * branch list list * (int*int*exn) list
   val normBr            : branch -> branch
   end;
@@ -1283,7 +1281,6 @@
       ((if !trace then warning ("blast: " ^ s) else ());
        Seq.empty);
 
-fun Blast_tac i = blast_tac (Data.claset()) i;
 
 
 (*** For debugging: these apply the prover to a subgoal and return
@@ -1294,25 +1291,24 @@
 (*Read a string to make an initial, singleton branch*)
 fun readGoal thy s = Syntax.read_prop_global thy s |> fromTerm thy |> rand |> mkGoal;
 
-fun tryInThy thy lim s =
+fun tryInThy thy cs lim s =
   let
     val state as State {fullTrace = ft, ...} = initialize thy;
     val res = timeap prove
-      (state, start_timing(), Data.claset(), [initBranch ([readGoal thy s], lim)], I);
+      (state, start_timing(), cs, [initBranch ([readGoal thy s], lim)], I);
     val _ = fullTrace := !ft;
   in res end;
 
 
 (** method setup **)
 
-val blast_method =
-  Args.bang_facts -- Scan.lift (Scan.option OuterParse.nat) --|
-  Method.sections Data.cla_modifiers >>
-  (fn (prems, NONE) => Data.cla_meth' blast_tac prems
-    | (prems, SOME lim) => Data.cla_meth' (fn cs => depth_tac cs lim) prems);
-
 val setup =
   setup_depth_limit #>
-  Method.setup @{binding blast} blast_method "tableau prover";
+  Method.setup @{binding blast}
+    (Args.bang_facts -- Scan.lift (Scan.option OuterParse.nat) --|
+      Method.sections Data.cla_modifiers >>
+      (fn (prems, NONE) => Data.cla_meth' blast_tac prems
+        | (prems, SOME lim) => Data.cla_meth' (fn cs => depth_tac cs lim) prems))
+    "classical tableau prover";
 
 end;
--- a/src/Provers/clasimp.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Provers/clasimp.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -26,8 +26,6 @@
   type clasimpset
   val get_css: Context.generic -> clasimpset
   val map_css: (clasimpset -> clasimpset) -> Context.generic -> Context.generic
-  val change_clasimpset: (clasimpset -> clasimpset) -> unit
-  val clasimpset: unit -> clasimpset
   val local_clasimpset_of: Proof.context -> clasimpset
   val addSIs2: clasimpset * thm list -> clasimpset
   val addSEs2: clasimpset * thm list -> clasimpset
@@ -42,19 +40,10 @@
   val addss': claset * simpset -> claset
   val addIffs: clasimpset * thm list -> clasimpset
   val delIffs: clasimpset * thm list -> clasimpset
-  val AddIffs: thm list -> unit
-  val DelIffs: thm list -> unit
-  val CLASIMPSET: (clasimpset -> tactic) -> tactic
-  val CLASIMPSET': (clasimpset -> 'a -> tactic) -> 'a -> tactic
   val clarsimp_tac: clasimpset -> int -> tactic
-  val Clarsimp_tac: int -> tactic
   val mk_auto_tac: clasimpset -> int -> int -> tactic
   val auto_tac: clasimpset -> tactic
-  val Auto_tac: tactic
-  val auto: unit -> unit
   val force_tac: clasimpset  -> int -> tactic
-  val Force_tac: int -> tactic
-  val force: int -> unit
   val fast_simp_tac: clasimpset -> int -> tactic
   val slow_simp_tac: clasimpset -> int -> tactic
   val best_simp_tac: clasimpset -> int -> tactic
@@ -84,9 +73,6 @@
   let val (cs', ss') = f (get_css context)
   in context |> Classical.map_cs (K cs') |> Simplifier.map_ss (K ss') end;
 
-fun change_clasimpset f = Context.>> (fn context => (Context.the_theory context; map_css f context));
-fun clasimpset () = (Classical.claset (), Simplifier.simpset ());
-
 fun local_clasimpset_of ctxt =
   (Classical.local_claset_of ctxt, Simplifier.local_simpset_of ctxt);
 
@@ -168,9 +154,6 @@
               Simplifier.addsimps);
 val op delIffs = Library.foldl (delIff Classical.delrules Simplifier.delsimps);
 
-fun AddIffs thms = change_clasimpset (fn css => css addIffs thms);
-fun DelIffs thms = change_clasimpset (fn css => css delIffs thms);
-
 fun addIffs_generic (x, ths) =
   Library.foldl (addIff (addXEs, addXIs) (addXEs, addXIs) #1) ((x, ()), ths) |> #1;
 
@@ -182,19 +165,10 @@
 
 (* tacticals *)
 
-fun CLASIMPSET tacf state =
-  Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss))) state;
-
-fun CLASIMPSET' tacf i state =
-  Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss) i)) state;
-
-
 fun clarsimp_tac (cs, ss) =
   safe_asm_full_simp_tac ss THEN_ALL_NEW
   Classical.clarify_tac (cs addSss ss);
 
-fun Clarsimp_tac i = clarsimp_tac (clasimpset ()) i;
-
 
 (* auto_tac *)
 
@@ -231,8 +205,6 @@
     end;
 
 fun auto_tac css = mk_auto_tac css 4 2;
-fun Auto_tac st = auto_tac (clasimpset ()) st;
-fun auto () = by Auto_tac;
 
 
 (* force_tac *)
@@ -242,8 +214,6 @@
         Classical.clarify_tac cs' 1,
         IF_UNSOLVED (Simplifier.asm_full_simp_tac ss 1),
         ALLGOALS (Classical.first_best_tac cs')]) end;
-fun Force_tac i = force_tac (clasimpset ()) i;
-fun force i = by (Force_tac i);
 
 
 (* basic combinations *)
--- a/src/Provers/classical.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Provers/classical.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -69,11 +69,7 @@
   val appSWrappers      : claset -> wrapper
   val appWrappers       : claset -> wrapper
 
-  val change_claset: (claset -> claset) -> unit
   val claset_of: theory -> claset
-  val claset: unit -> claset
-  val CLASET: (claset -> tactic) -> tactic
-  val CLASET': (claset -> 'a -> tactic) -> 'a -> tactic
   val local_claset_of   : Proof.context -> claset
 
   val fast_tac          : claset -> int -> tactic
@@ -107,24 +103,6 @@
   val inst_step_tac     : claset -> int -> tactic
   val inst0_step_tac    : claset -> int -> tactic
   val instp_step_tac    : claset -> int -> tactic
-
-  val AddDs             : thm list -> unit
-  val AddEs             : thm list -> unit
-  val AddIs             : thm list -> unit
-  val AddSDs            : thm list -> unit
-  val AddSEs            : thm list -> unit
-  val AddSIs            : thm list -> unit
-  val Delrules          : thm list -> unit
-  val Safe_tac          : tactic
-  val Safe_step_tac     : int -> tactic
-  val Clarify_tac       : int -> tactic
-  val Clarify_step_tac  : int -> tactic
-  val Step_tac          : int -> tactic
-  val Fast_tac          : int -> tactic
-  val Best_tac          : int -> tactic
-  val Slow_tac          : int -> tactic
-  val Slow_best_tac     : int -> tactic
-  val Deepen_tac        : int -> int -> tactic
 end;
 
 signature CLASSICAL =
@@ -870,23 +848,9 @@
 fun map_context_cs f = GlobalClaset.map (apsnd
   (fn ContextCS {swrappers, uwrappers} => make_context_cs (f (swrappers, uwrappers))));
 
-fun change_claset f = Context.>> (Context.map_theory (map_claset f));
-
 fun claset_of thy =
   let val (cs, ctxt_cs) = GlobalClaset.get thy
   in context_cs (ProofContext.init thy) cs (ctxt_cs) end;
-val claset = claset_of o ML_Context.the_global_context;
-
-fun CLASET tacf st = tacf (claset_of (Thm.theory_of_thm st)) st;
-fun CLASET' tacf i st = tacf (claset_of (Thm.theory_of_thm st)) i st;
-
-fun AddDs args = change_claset (fn cs => cs addDs args);
-fun AddEs args = change_claset (fn cs => cs addEs args);
-fun AddIs args = change_claset (fn cs => cs addIs args);
-fun AddSDs args = change_claset (fn cs => cs addSDs args);
-fun AddSEs args = change_claset (fn cs => cs addSEs args);
-fun AddSIs args = change_claset (fn cs => cs addSIs args);
-fun Delrules args = change_claset (fn cs => cs delrules args);
 
 
 (* context dependent components *)
@@ -930,21 +894,6 @@
 val rule_del = attrib delrule o ContextRules.rule_del;
 
 
-(* tactics referring to the implicit claset *)
-
-(*the abstraction over the proof state delays the dereferencing*)
-fun Safe_tac st           = safe_tac (claset()) st;
-fun Safe_step_tac i st    = safe_step_tac (claset()) i st;
-fun Clarify_step_tac i st = clarify_step_tac (claset()) i st;
-fun Clarify_tac i st      = clarify_tac (claset()) i st;
-fun Step_tac i st         = step_tac (claset()) i st;
-fun Fast_tac i st         = fast_tac (claset()) i st;
-fun Best_tac i st         = best_tac (claset()) i st;
-fun Slow_tac i st         = slow_tac (claset()) i st;
-fun Slow_best_tac i st    = slow_best_tac (claset()) i st;
-fun Deepen_tac m          = deepen_tac (claset()) m;
-
-
 end;
 
 
@@ -972,15 +921,12 @@
 
 (** proof methods **)
 
-fun METHOD_CLASET tac ctxt = METHOD (tac ctxt (local_claset_of ctxt));
-fun METHOD_CLASET' tac ctxt = METHOD (HEADGOAL o tac ctxt (local_claset_of ctxt));
-
-
 local
 
-fun some_rule_tac ctxt (CS {xtra_netpair, ...}) facts = SUBGOAL (fn (goal, i) =>
+fun some_rule_tac ctxt facts = SUBGOAL (fn (goal, i) =>
   let
     val [rules1, rules2, rules4] = ContextRules.find_rules false facts goal ctxt;
+    val CS {xtra_netpair, ...} = local_claset_of ctxt;
     val rules3 = ContextRules.find_rules_netpair true facts goal xtra_netpair;
     val rules = rules1 @ rules2 @ rules3 @ rules4;
     val ruleq = Drule.multi_resolves facts rules;
@@ -990,16 +936,15 @@
   end)
   THEN_ALL_NEW Goal.norm_hhf_tac;
 
-fun rule_tac [] ctxt cs facts = some_rule_tac ctxt cs facts
-  | rule_tac rules _ _ facts = Method.rule_tac rules facts;
+in
 
-fun default_tac rules ctxt cs facts =
-  HEADGOAL (rule_tac rules ctxt cs facts) ORELSE
+fun rule_tac ctxt [] facts = some_rule_tac ctxt facts
+  | rule_tac _ rules facts = Method.rule_tac rules facts;
+
+fun default_tac ctxt rules facts =
+  HEADGOAL (rule_tac ctxt rules facts) ORELSE
   Class.default_intro_tac ctxt facts;
 
-in
-  val rule = METHOD_CLASET' o rule_tac;
-  val default = METHOD_CLASET o default_tac;
 end;
 
 
@@ -1033,9 +978,11 @@
 (** setup_methods **)
 
 val setup_methods =
-  Method.setup @{binding default} (Attrib.thms >> default)
+  Method.setup @{binding default}
+   (Attrib.thms >> (fn rules => fn ctxt => METHOD (default_tac ctxt rules)))
     "apply some intro/elim rule (potentially classical)" #>
-  Method.setup @{binding rule} (Attrib.thms >> rule)
+  Method.setup @{binding rule}
+    (Attrib.thms >> (fn rules => fn ctxt => METHOD (HEADGOAL o rule_tac ctxt rules)))
     "apply some intro/elim rule (potentially classical)" #>
   Method.setup @{binding contradiction} (Scan.succeed (K contradiction))
     "proof by contradiction" #>
--- a/src/Provers/splitter.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Provers/splitter.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -35,8 +35,6 @@
   val split_asm_tac   : thm list -> int -> tactic
   val addsplits       : simpset * thm list -> simpset
   val delsplits       : simpset * thm list -> simpset
-  val Addsplits       : thm list -> unit
-  val Delsplits       : thm list -> unit
   val split_add: attribute
   val split_del: attribute
   val split_modifiers : Method.modifier parser list
@@ -453,9 +451,6 @@
         in Simplifier.delloop (ss, split_name name asm)
   end in Library.foldl delsplit (ss,splits) end;
 
-fun Addsplits splits = (change_simpset (fn ss => ss addsplits splits));
-fun Delsplits splits = (change_simpset (fn ss => ss delsplits splits));
-
 
 (* attributes *)
 
@@ -472,14 +467,13 @@
   Args.$$$ splitN -- Args.add -- Args.colon >> K (I, split_add),
   Args.$$$ splitN -- Args.del -- Args.colon >> K (I, split_del)];
 
-fun split_meth parser = (Attrib.thms >>
-  (fn ths => K (SIMPLE_METHOD' (CHANGED_PROP o gen_split_tac ths)))) parser;
-
 
 (* theory setup *)
 
 val setup =
   Attrib.setup @{binding split} (Attrib.add_del split_add split_del) "declare case split rule" #>
-  Method.setup @{binding split} split_meth "apply case split rule";
+  Method.setup @{binding split}
+    (Attrib.thms >> (fn ths => K (SIMPLE_METHOD' (CHANGED_PROP o gen_split_tac ths))))
+    "apply case split rule";
 
 end;
--- a/src/Pure/Concurrent/future.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Concurrent/future.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -44,6 +44,7 @@
   val join_result: 'a future -> 'a Exn.result
   val join: 'a future -> 'a
   val map: ('a -> 'b) -> 'a future -> 'b future
+  val interruptible_task: ('a -> 'b) -> 'a -> 'b
   val interrupt_task: string -> unit
   val cancel_group: group -> unit
   val cancel: 'a future -> unit
@@ -211,7 +212,8 @@
     val _ = if continue then () else scheduler := NONE;
 
     val _ = notify_all ();
-    val _ = interruptible (fn () => wait_timeout (Time.fromSeconds 1)) ()
+    val _ = interruptible (fn () =>
+        wait_timeout (Time.fromMilliseconds (if null (! canceled) then 1000 else 50))) ()
       handle Exn.Interrupt => List.app do_cancel (Task_Queue.cancel_all (! queue));
   in continue end;
 
@@ -236,7 +238,7 @@
 fun future_job group (e: unit -> 'a) =
   let
     val result = ref (NONE: 'a Exn.result option);
-    val job = Multithreading.with_attributes (Thread.getAttributes ())
+    val job = Multithreading.with_attributes Multithreading.restricted_interrupts
       (fn _ => fn ok =>
         let
           val res = if ok then Exn.capture e () else Exn.Exn Exn.Interrupt;
@@ -350,6 +352,15 @@
 
 (* cancellation *)
 
+fun interruptible_task f x =
+  if Multithreading.available then
+    Multithreading.with_attributes
+      (if is_some (thread_data ())
+       then Multithreading.restricted_interrupts
+       else Multithreading.regular_interrupts)
+      (fn _ => f) x
+  else interruptible f x;
+
 (*interrupt: permissive signal, may get ignored*)
 fun interrupt_task id = SYNCHRONIZED "interrupt"
   (fn () => Task_Queue.interrupt_external (! queue) id);
--- a/src/Pure/General/ROOT.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/General/ROOT.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -15,6 +15,7 @@
 use "seq.ML";
 use "position.ML";
 use "symbol_pos.ML";
+use "antiquote.ML";
 use "../ML/ml_lex.ML";
 use "../ML/ml_parse.ML";
 use "secure.ML";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/General/antiquote.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -0,0 +1,102 @@
+(*  Title:      Pure/General/antiquote.ML
+    Author:     Markus Wenzel, TU Muenchen
+
+Text with antiquotations of inner items (types, terms, theorems etc.).
+*)
+
+signature ANTIQUOTE =
+sig
+  datatype 'a antiquote =
+    Text of 'a |
+    Antiq of Symbol_Pos.T list * Position.range |
+    Open of Position.T |
+    Close of Position.T
+  val is_text: 'a antiquote -> bool
+  val report: ('a -> unit) -> 'a antiquote -> unit
+  val check_nesting: 'a antiquote list -> unit
+  val scan: Symbol_Pos.T list -> 'a antiquote * Symbol_Pos.T list
+  val scan_text: Symbol_Pos.T list -> Symbol_Pos.T list antiquote * Symbol_Pos.T list
+  val read: Symbol_Pos.T list * Position.T -> Symbol_Pos.T list antiquote list
+end;
+
+structure Antiquote: ANTIQUOTE =
+struct
+
+(* datatype antiquote *)
+
+datatype 'a antiquote =
+  Text of 'a |
+  Antiq of Symbol_Pos.T list * Position.range |
+  Open of Position.T |
+  Close of Position.T;
+
+fun is_text (Text _) = true
+  | is_text _ = false;
+
+
+(* report *)
+
+val report_antiq = Position.report Markup.antiq;
+
+fun report report_text (Text x) = report_text x
+  | report _ (Antiq (_, (pos, _))) = report_antiq pos
+  | report _ (Open pos) = report_antiq pos
+  | report _ (Close pos) = report_antiq pos;
+
+
+(* check_nesting *)
+
+fun err_unbalanced pos =
+  error ("Unbalanced antiquotation block parentheses" ^ Position.str_of pos);
+
+fun check_nesting antiqs =
+  let
+    fun check [] [] = ()
+      | check [] (pos :: _) = err_unbalanced pos
+      | check (Open pos :: ants) ps = check ants (pos :: ps)
+      | check (Close pos :: _) [] = err_unbalanced pos
+      | check (Close _ :: ants) (_ :: ps) = check ants ps
+      | check (_ :: ants) ps = check ants ps;
+  in check antiqs [] end;
+
+
+(* scan *)
+
+open Basic_Symbol_Pos;
+
+local
+
+val scan_txt =
+  $$$ "@" --| Scan.ahead (~$$$ "{") ||
+  Scan.one (fn (s, _) => s <> "@" andalso s <> "\\<lbrace>" andalso s <> "\\<rbrace>"
+    andalso Symbol.is_regular s) >> single;
+
+val scan_ant =
+  Symbol_Pos.scan_quoted ||
+  Scan.one (fn (s, _) => s <> "}" andalso Symbol.is_regular s) >> single;
+
+val scan_antiq =
+  Symbol_Pos.scan_pos -- ($$$ "@" |-- $$$ "{" |--
+    Symbol_Pos.!!! "missing closing brace of antiquotation"
+      (Scan.repeat scan_ant -- ($$$ "}" |-- Symbol_Pos.scan_pos)))
+  >> (fn (pos1, (body, pos2)) => (flat body, Position.range pos1 pos2));
+
+val scan_open = Symbol_Pos.scan_pos --| $$$ "\\<lbrace>";
+val scan_close = Symbol_Pos.scan_pos --| $$$ "\\<rbrace>";
+
+in
+
+fun scan x = (scan_antiq >> Antiq || scan_open >> Open || scan_close >> Close) x;
+val scan_text = scan || Scan.repeat1 scan_txt >> (Text o flat);
+
+end;
+
+
+(* read *)
+
+fun read (syms, pos) =
+  (case Scan.read Symbol_Pos.stopper (Scan.repeat scan_text) syms of
+    SOME xs => (List.app (report (K ())) xs; check_nesting xs; xs)
+  | NONE => error ("Malformed quotation/antiquotation source" ^ Position.str_of pos));
+
+end;
--- a/src/Pure/General/markup.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/General/markup.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -62,6 +62,17 @@
   val propN: string val prop: T
   val attributeN: string val attribute: string -> T
   val methodN: string val method: string -> T
+  val ML_keywordN: string val ML_keyword: T
+  val ML_identN: string val ML_ident: T
+  val ML_tvarN: string val ML_tvar: T
+  val ML_numeralN: string val ML_numeral: T
+  val ML_charN: string val ML_char: T
+  val ML_stringN: string val ML_string: T
+  val ML_commentN: string val ML_comment: T
+  val ML_malformedN: string val ML_malformed: T
+  val ML_defN: string val ML_def: T
+  val ML_refN: string val ML_ref: T
+  val ML_typingN: string val ML_typing: T
   val ML_sourceN: string val ML_source: T
   val doc_sourceN: string val doc_source: T
   val antiqN: string val antiq: T
@@ -213,6 +224,22 @@
 val (methodN, method) = markup_string "method" nameN;
 
 
+(* ML syntax *)
+
+val (ML_keywordN, ML_keyword) = markup_elem "ML_keyword";
+val (ML_identN, ML_ident) = markup_elem "ML_ident";
+val (ML_tvarN, ML_tvar) = markup_elem "ML_tvar";
+val (ML_numeralN, ML_numeral) = markup_elem "ML_numeral";
+val (ML_charN, ML_char) = markup_elem "ML_char";
+val (ML_stringN, ML_string) = markup_elem "ML_string";
+val (ML_commentN, ML_comment) = markup_elem "ML_comment";
+val (ML_malformedN, ML_malformed) = markup_elem "ML_malformed";
+
+val (ML_defN, ML_def) = markup_elem "ML_def";
+val (ML_refN, ML_ref) = markup_elem "ML_ref";
+val (ML_typingN, ML_typing) = markup_elem "ML_typing";
+
+
 (* embedded source text *)
 
 val (ML_sourceN, ML_source) = markup_elem "ML_source";
--- a/src/Pure/General/markup.scala	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/General/markup.scala	Sat Mar 28 00:13:01 2009 +0100
@@ -80,6 +80,22 @@
   val DOC_ANTIQ = "doc_antiq"
 
 
+  /* ML syntax */
+
+  val ML_KEYWORD = "ML_keyword"
+  val ML_IDENT = "ML_ident"
+  val ML_TVAR = "ML_tvar"
+  val ML_NUMERAL = "ML_numeral"
+  val ML_CHAR = "ML_char"
+  val ML_STRING = "ML_string"
+  val ML_COMMENT = "ML_comment"
+  val ML_MALFORMED = "ML_malformed"
+
+  val ML_DEF = "ML_def"
+  val ML_REF = "ML_ref"
+  val ML_TYPING = "ML_typing"
+
+
   /* outer syntax */
 
   val KEYWORD_DECL = "keyword_decl"
--- a/src/Pure/General/output.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/General/output.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -47,7 +47,6 @@
   val debugging: bool ref
   val no_warnings: ('a -> 'b) -> 'a -> 'b
   val debug: (unit -> string) -> unit
-  val ml_output: (string -> unit) * (string -> 'a)
   val accumulated_time: unit -> unit
 end;
 
@@ -101,7 +100,7 @@
 val error_fn = ref (std_output o suffix "\n" o prefix_lines "*** ");
 val debug_fn = ref (std_output o suffix "\n" o prefix_lines "::: ");
 val prompt_fn = ref std_output;
-val status_fn = ref (fn s => ! writeln_fn s);
+val status_fn = ref (fn _: string => ());
 
 fun writeln s = ! writeln_fn (output s);
 fun priority s = ! priority_fn (output s);
@@ -120,8 +119,6 @@
 val debugging = ref false;
 fun debug s = if ! debugging then ! debug_fn (output (s ())) else ()
 
-val ml_output = (writeln, error);
-
 
 
 (** timing **)
--- a/src/Pure/General/position.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/General/position.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -24,6 +24,7 @@
   val of_properties: Properties.T -> T
   val properties_of: T -> Properties.T
   val default_properties: T -> Properties.T -> Properties.T
+  val report_text: Markup.T -> T -> string -> unit
   val report: Markup.T -> T -> unit
   val str_of: T -> string
   type range = T * T
@@ -121,9 +122,11 @@
   if exists (member (op =) Markup.position_properties o #1) props then props
   else properties_of default @ props;
 
-fun report markup (pos as Pos (count, _)) =
+fun report_text markup (pos as Pos (count, _)) txt =
   if invalid_count count then ()
-  else Output.status (Markup.markup (Markup.properties (properties_of pos) markup) "");
+  else Output.status (Markup.markup (Markup.properties (properties_of pos) markup) txt);
+
+fun report markup pos = report_text markup pos "";
 
 
 (* str_of *)
--- a/src/Pure/General/pretty.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/General/pretty.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -19,9 +19,6 @@
 a unit for breaking).
 *)
 
-type pprint_args = (output -> unit) * (int -> unit) * (int -> unit) *
-  (unit -> unit) * (unit -> unit);
-
 signature PRETTY =
 sig
   val default_indent: string -> int -> output
@@ -49,6 +46,7 @@
   val commas: T list -> T list
   val enclose: string -> string -> T list -> T
   val enum: string -> string -> string -> T list -> T
+  val position: Position.T -> T
   val list: string -> string -> T list -> T
   val str_list: string -> string -> string list -> T
   val big_list: string -> T list -> T
@@ -57,7 +55,8 @@
   val setmargin: int -> unit
   val setmp_margin: int -> ('a -> 'b) -> 'a -> 'b
   val setdepth: int -> unit
-  val pprint: T -> pprint_args -> unit
+  val to_ML: T -> ML_Pretty.pretty
+  val from_ML: ML_Pretty.pretty -> T
   val symbolicN: string
   val output_buffer: T -> Buffer.T
   val output: T -> output
@@ -105,9 +104,9 @@
 (** printing items: compound phrases, strings, and breaks **)
 
 datatype T =
-  Block of Markup.T * T list * int * int |  (*markup, body, indentation, length*)
-  String of output * int |                  (*text, length*)
-  Break of bool * int;                      (*mandatory flag, width if not taken*)
+  Block of (output * output) * T list * int * int |  (*markup output, body, indentation, length*)
+  String of output * int |                           (*text, length*)
+  Break of bool * int;                               (*mandatory flag, width if not taken*)
 
 fun length (Block (_, _, _, len)) = len
   | length (String (_, len)) = len
@@ -125,12 +124,14 @@
 fun breaks prts = Library.separate (brk 1) prts;
 fun fbreaks prts = Library.separate fbrk prts;
 
-fun markup_block m (indent, es) =
+fun block_markup m (indent, es) =
   let
     fun sum [] k = k
       | sum (e :: es) k = sum es (length e + k);
   in Block (m, es, indent, sum es 0) end;
 
+fun markup_block m arg = block_markup (Markup.output m) arg;
+
 val blk = markup_block Markup.none;
 fun block prts = blk (2, prts);
 val strs = block o breaks o map str;
@@ -159,6 +160,9 @@
 
 fun enum sep lpar rpar prts = enclose lpar rpar (separate sep prts);
 
+val position =
+  enum "," "{" "}" o map (fn (x, y) => str (x ^ "=" ^ y)) o Position.properties_of;
+
 val list = enum ",";
 fun str_list lpar rpar strs = list lpar rpar (map str strs);
 
@@ -195,7 +199,7 @@
 local
   fun pruning dp (Block (m, bes, indent, wd)) =
         if dp > 0
-        then markup_block m (indent, map (pruning (dp - 1)) bes)
+        then block_markup m (indent, map (pruning (dp - 1)) bes)
         else str "..."
     | pruning dp e = e
 in
@@ -261,7 +265,7 @@
 fun format ([], _, _) text = text
   | format (e :: es, block as (blockind, blockin), after) (text as {ind, pos, nl, ...}) =
       (case e of
-        Block (markup, bes, indent, wd) =>
+        Block ((bg, en), bes, indent, wd) =>
           let
             val {emergencypos, ...} = ! margin_info;
             val pos' = pos + indent;
@@ -269,7 +273,6 @@
             val block' =
               if pos' < emergencypos then (ind |> add_indent indent, pos')
               else (add_indent pos'' Buffer.empty, pos'');
-            val (bg, en) = Markup.output markup;
             val btext: text = text
               |> control bg
               |> format (bes, block', breakdist (es, after))
@@ -301,9 +304,9 @@
 (*symbolic markup -- no formatting*)
 fun symbolic prt =
   let
-    fun out (Block (m, [], _, _)) = Buffer.markup m I
-      | out (Block (m, prts, indent, _)) =
-          Buffer.markup m (Buffer.markup (Markup.block indent) (fold out prts))
+    fun out (Block ((bg, en), [], _, _)) = Buffer.add bg #> Buffer.add en
+      | out (Block ((bg, en), prts, indent, _)) =
+          Buffer.add bg #> Buffer.markup (Markup.block indent) (fold out prts) #> Buffer.add en
       | out (String (s, _)) = Buffer.add s
       | out (Break (false, wd)) = Buffer.markup (Markup.break wd) (Buffer.add (output_spaces wd))
       | out (Break (true, _)) = Buffer.markup Markup.fbreak (Buffer.add (output_spaces 1));
@@ -312,26 +315,22 @@
 (*unformatted output*)
 fun unformatted prt =
   let
-    fun fmt (Block (m, prts, _, _)) = Buffer.markup m (fold fmt prts)
+    fun fmt (Block ((bg, en), prts, _, _)) = Buffer.add bg #> fold fmt prts #> Buffer.add en
       | fmt (String (s, _)) = Buffer.add s
       | fmt (Break (false, wd)) = Buffer.add (output_spaces wd)
       | fmt (Break (true, _)) = Buffer.add (output_spaces 1);
   in fmt (prune prt) Buffer.empty end;
 
-(*ML toplevel pretty printing*)
-fun pprint prt (put_str0, begin_blk, put_brk, put_fbrk, end_blk) =
-  let
-    fun put_str "" = ()
-      | put_str s = put_str0 s;
-    fun pp (Block (m, prts, ind, _)) =
-          let val (bg, en) = Markup.output m
-          in put_str bg; begin_blk ind; pp_lst prts; end_blk (); put_str en end
-      | pp (String (s, _)) = put_str s
-      | pp (Break (false, wd)) = put_brk wd
-      | pp (Break (true, _)) = put_fbrk ()
-    and pp_lst [] = ()
-      | pp_lst (prt :: prts) = (pp prt; pp_lst prts);
-  in pp (prune prt) end;
+
+(* ML toplevel pretty printing *)
+
+fun to_ML (Block (m, prts, ind, _)) = ML_Pretty.Block (m, map to_ML prts, ind)
+  | to_ML (String s) = ML_Pretty.String s
+  | to_ML (Break b) = ML_Pretty.Break b;
+
+fun from_ML (ML_Pretty.Block (m, prts, ind)) = block_markup m (ind, map from_ML prts)
+  | from_ML (ML_Pretty.String s) = String s
+  | from_ML (ML_Pretty.Break b) = Break b;
 
 
 (* output interfaces *)
--- a/src/Pure/General/secure.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/General/secure.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -9,11 +9,9 @@
   val set_secure: unit -> unit
   val is_secure: unit -> bool
   val deny_secure: string -> unit
-  val use_text: ML_NameSpace.nameSpace -> int * string ->
-    (string -> unit) * (string -> 'a) -> bool -> string -> unit
-  val use_file: ML_NameSpace.nameSpace ->
-    (string -> unit) * (string -> 'a) -> bool -> string -> unit
-  val use: string -> unit
+  val use_text: use_context -> int * string -> bool -> string -> unit
+  val use_file: use_context -> bool -> string -> unit
+  val toplevel_pp: string list -> string -> unit
   val commit: unit -> unit
   val system_out: string -> string * int
   val system: string -> int
@@ -39,19 +37,17 @@
 
 fun secure_mltext () = deny_secure "Cannot evaluate ML source in secure mode";
 
-fun raw_use_text ns = use_text ML_Parse.fix_ints (Position.str_of oo Position.line_file) ns;
-fun raw_use_file ns = use_file ML_Parse.fix_ints (Position.str_of oo Position.line_file) ns;
+val raw_use_text = use_text;
+val raw_use_file = use_file;
+val raw_toplevel_pp = toplevel_pp;
 
-fun use_text ns pos pr verbose txt =
-  (secure_mltext (); raw_use_text ns pos pr verbose txt);
+fun use_text context pos verbose txt = (secure_mltext (); raw_use_text context pos verbose txt);
+fun use_file context verbose name = (secure_mltext (); raw_use_file context verbose name);
 
-fun use_file ns pr verbose name =
-  (secure_mltext (); raw_use_file ns pr verbose name);
-
-fun use name = use_file ML_NameSpace.global Output.ml_output true name;
+fun toplevel_pp path pp = (secure_mltext (); raw_toplevel_pp ML_Parse.global_context path pp);
 
 (*commit is dynamically bound!*)
-fun commit () = raw_use_text ML_NameSpace.global (0, "") Output.ml_output false "commit();";
+fun commit () = raw_use_text ML_Parse.global_context (0, "") false "commit();";
 
 
 (* shell commands *)
@@ -72,6 +68,8 @@
 (*override previous toplevel bindings!*)
 val use_text = Secure.use_text;
 val use_file = Secure.use_file;
-fun use s = Secure.use s handle ERROR msg => (writeln msg; raise Fail "ML error");
+fun use s = Secure.use_file ML_Parse.global_context true s
+  handle ERROR msg => (writeln msg; raise Fail "ML error");
+val toplevel_pp = Secure.toplevel_pp;
 val system_out = Secure.system_out;
 val system = Secure.system;
--- a/src/Pure/General/symbol_pos.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/General/symbol_pos.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -20,7 +20,11 @@
   val is_eof: T -> bool
   val stopper: T Scan.stopper
   val !!! : string -> (T list -> 'a) -> T list -> 'a
+  val change_prompt: ('a -> 'b) -> 'a -> 'b
   val scan_pos: T list -> Position.T * T list
+  val scan_string: T list -> (Position.T * (T list * Position.T)) * T list
+  val scan_alt_string: T list -> (Position.T * (T list * Position.T)) * T list
+  val scan_quoted: T list -> T list * T list
   val scan_comment: (string -> (T list -> T list * T list) -> T list -> T list * T list) ->
     T list -> T list * T list
   val scan_comment_body: (string -> (T list -> T list * T list) -> T list -> T list * T list) ->
@@ -34,7 +38,7 @@
   val explode: text * Position.T -> T list
 end;
 
-structure SymbolPos: SYMBOL_POS =
+structure Symbol_Pos: SYMBOL_POS =
 struct
 
 (* type T *)
@@ -83,12 +87,44 @@
       (case msg of NONE => "" | SOME s => "\n" ^ s);
   in Scan.!! err scan end;
 
+fun change_prompt scan = Scan.prompt "# " scan;
+
 fun $$$ s = Scan.one (fn x => symbol x = s) >> single;
 fun ~$$$ s = Scan.one (fn x => symbol x <> s) >> single;
 
 val scan_pos = Scan.ahead (Scan.one (K true)) >> (fn (_, pos): T => pos);
 
 
+(* Isabelle strings *)
+
+local
+
+val char_code =
+  Scan.one (Symbol.is_ascii_digit o symbol) --
+  Scan.one (Symbol.is_ascii_digit o symbol) --
+  Scan.one (Symbol.is_ascii_digit o symbol) :|--
+  (fn (((a, pos), (b, _)), (c, _)) =>
+    let val (n, _) = Library.read_int [a, b, c]
+    in if n <= 255 then Scan.succeed [(chr n, pos)] else Scan.fail end);
+
+fun scan_str q =
+  $$$ "\\" |-- !!! "bad escape character in string" ($$$ q || $$$ "\\" || char_code) ||
+  Scan.one (fn (s, _) => s <> q andalso s <> "\\" andalso Symbol.is_regular s) >> single;
+
+fun scan_strs q =
+  (scan_pos --| $$$ q) -- !!! "missing quote at end of string"
+    (change_prompt ((Scan.repeat (scan_str q) >> flat) -- ($$$ q |-- scan_pos)));
+
+in
+
+val scan_string = scan_strs "\"";
+val scan_alt_string = scan_strs "`";
+
+val scan_quoted = Scan.trace (scan_string || scan_alt_string) >> #2;
+
+end;
+
+
 (* ML-style comments *)
 
 local
@@ -99,7 +135,7 @@
   Scan.lift ($$$ "*" --| Scan.ahead (~$$$ ")")) ||
   Scan.lift (Scan.one (fn (s, _) => s <> "*" andalso Symbol.is_regular s)) >> single;
 
-val scan_body = Scan.pass 0 (Scan.repeat scan_cmt >> flat);
+val scan_body = change_prompt (Scan.pass 0 (Scan.repeat scan_cmt >> flat));
 
 in
 
@@ -149,5 +185,5 @@
 
 end;
 
-structure BasicSymbolPos: BASIC_SYMBOL_POS = SymbolPos;   (*not open by default*)
+structure Basic_Symbol_Pos: BASIC_SYMBOL_POS = Symbol_Pos;   (*not open by default*)
 
--- a/src/Pure/General/table.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/General/table.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -383,3 +383,6 @@
 
 structure Inttab = TableFun(type key = int val ord = int_ord);
 structure Symtab = TableFun(type key = string val ord = fast_string_ord);
+structure Symreltab = TableFun(type key = string * string
+  val ord = prod_ord fast_string_ord fast_string_ord);
+
--- a/src/Pure/IsaMakefile	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/IsaMakefile	Sat Mar 28 00:13:01 2009 +0100
@@ -20,17 +20,17 @@
 ## Pure
 
 BOOTSTRAP_FILES = ML-Systems/exn.ML ML-Systems/ml_name_space.ML		\
-  ML-Systems/mosml.ML ML-Systems/multithreading.ML			\
-  ML-Systems/multithreading_polyml.ML ML-Systems/overloading_smlnj.ML	\
-  ML-Systems/polyml-4.1.3.ML ML-Systems/polyml-4.1.4.ML			\
-  ML-Systems/polyml-4.2.0.ML ML-Systems/polyml-5.0.ML			\
-  ML-Systems/polyml-5.1.ML ML-Systems/polyml-experimental.ML		\
-  ML-Systems/polyml.ML ML-Systems/polyml_common.ML			\
-  ML-Systems/polyml_old_compiler4.ML					\
-  ML-Systems/polyml_old_compiler5.ML ML-Systems/proper_int.ML		\
-  ML-Systems/smlnj.ML ML-Systems/system_shell.ML			\
-  ML-Systems/thread_dummy.ML ML-Systems/time_limit.ML			\
-  ML-Systems/universal.ML
+  ML-Systems/ml_pretty.ML ML-Systems/mosml.ML				\
+  ML-Systems/multithreading.ML ML-Systems/multithreading_polyml.ML	\
+  ML-Systems/overloading_smlnj.ML ML-Systems/polyml-4.1.3.ML		\
+  ML-Systems/polyml-4.1.4.ML ML-Systems/polyml-4.2.0.ML			\
+  ML-Systems/polyml-5.0.ML ML-Systems/polyml-5.1.ML			\
+  ML-Systems/polyml-experimental.ML ML-Systems/polyml.ML		\
+  ML-Systems/polyml_common.ML ML-Systems/polyml_old_compiler4.ML	\
+  ML-Systems/polyml_old_compiler5.ML ML-Systems/polyml_pp.ML		\
+  ML-Systems/proper_int.ML ML-Systems/smlnj.ML				\
+  ML-Systems/system_shell.ML ML-Systems/thread_dummy.ML			\
+  ML-Systems/time_limit.ML ML-Systems/universal.ML
 
 RAW: $(OUT)/RAW
 
@@ -45,17 +45,17 @@
   Concurrent/mailbox.ML Concurrent/par_list.ML				\
   Concurrent/par_list_dummy.ML Concurrent/simple_thread.ML		\
   Concurrent/synchronized.ML Concurrent/task_queue.ML General/ROOT.ML	\
-  General/alist.ML General/balanced_tree.ML General/basics.ML		\
-  General/binding.ML General/buffer.ML General/file.ML			\
-  General/graph.ML General/heap.ML General/integer.ML General/lazy.ML	\
-  General/long_name.ML General/markup.ML General/name_space.ML		\
-  General/ord_list.ML General/output.ML General/path.ML			\
-  General/position.ML General/pretty.ML General/print_mode.ML		\
-  General/properties.ML General/queue.ML General/scan.ML		\
-  General/secure.ML General/seq.ML General/source.ML General/stack.ML	\
-  General/symbol.ML General/symbol_pos.ML General/table.ML		\
-  General/url.ML General/xml.ML General/yxml.ML Isar/ROOT.ML		\
-  Isar/antiquote.ML Isar/args.ML Isar/attrib.ML Isar/auto_bind.ML	\
+  General/alist.ML General/antiquote.ML General/balanced_tree.ML	\
+  General/basics.ML General/binding.ML General/buffer.ML		\
+  General/file.ML General/graph.ML General/heap.ML General/integer.ML	\
+  General/lazy.ML General/long_name.ML General/markup.ML		\
+  General/name_space.ML General/ord_list.ML General/output.ML		\
+  General/path.ML General/position.ML General/pretty.ML			\
+  General/print_mode.ML General/properties.ML General/queue.ML		\
+  General/scan.ML General/secure.ML General/seq.ML General/source.ML	\
+  General/stack.ML General/symbol.ML General/symbol_pos.ML		\
+  General/table.ML General/url.ML General/xml.ML General/yxml.ML	\
+  Isar/ROOT.ML Isar/args.ML Isar/attrib.ML Isar/auto_bind.ML		\
   Isar/calculation.ML Isar/class.ML Isar/class_target.ML Isar/code.ML	\
   Isar/code_unit.ML Isar/constdefs.ML Isar/context_rules.ML		\
   Isar/element.ML Isar/expression.ML Isar/isar_cmd.ML			\
@@ -69,7 +69,9 @@
   Isar/spec_parse.ML Isar/specification.ML Isar/theory_target.ML	\
   Isar/toplevel.ML Isar/value_parse.ML ML/ml_antiquote.ML		\
   ML/ml_context.ML ML/ml_lex.ML ML/ml_parse.ML ML/ml_syntax.ML		\
-  ML/ml_thms.ML ML-Systems/install_pp_polyml.ML Proof/extraction.ML	\
+  ML/ml_test.ML ML/ml_thms.ML ML-Systems/install_pp_polyml.ML		\
+  ML-Systems/install_pp_polyml-experimental.ML				\
+  ML-Systems/use_context.ML Proof/extraction.ML				\
   Proof/proof_rewrite_rules.ML Proof/proof_syntax.ML			\
   Proof/proofchecker.ML Proof/reconstruct.ML ProofGeneral/ROOT.ML	\
   ProofGeneral/pgip.ML ProofGeneral/pgip_input.ML			\
--- a/src/Pure/Isar/ROOT.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/ROOT.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -24,7 +24,6 @@
 use "outer_parse.ML";
 use "value_parse.ML";
 use "args.ML";
-use "antiquote.ML";
 use "../ML/ml_context.ML";
 
 (*theory sources*)
--- a/src/Pure/Isar/antiquote.ML	Sat Mar 28 00:11:02 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,109 +0,0 @@
-(*  Title:      Pure/Isar/antiquote.ML
-    Author:     Markus Wenzel, TU Muenchen
-
-Text with antiquotations of inner items (terms, types etc.).
-*)
-
-signature ANTIQUOTE =
-sig
-  datatype antiquote =
-    Text of string | Antiq of SymbolPos.T list * Position.T |
-    Open of Position.T | Close of Position.T
-  val is_antiq: antiquote -> bool
-  val read: SymbolPos.T list * Position.T -> antiquote list
-  val read_antiq: Scan.lexicon -> (OuterLex.token list -> 'a * OuterLex.token list) ->
-    SymbolPos.T list * Position.T -> 'a
-end;
-
-structure Antiquote: ANTIQUOTE =
-struct
-
-(* datatype antiquote *)
-
-datatype antiquote =
-  Text of string |
-  Antiq of SymbolPos.T list * Position.T |
-  Open of Position.T |
-  Close of Position.T;
-
-fun is_antiq (Text _) = false
-  | is_antiq _ = true;
-
-
-(* check_nesting *)
-
-fun err_unbalanced pos =
-  error ("Unbalanced antiquotation block parentheses" ^ Position.str_of pos);
-
-fun check_nesting antiqs =
-  let
-    fun check [] [] = ()
-      | check [] (pos :: _) = err_unbalanced pos
-      | check (Open pos :: ants) ps = check ants (pos :: ps)
-      | check (Close pos :: _) [] = err_unbalanced pos
-      | check (Close _ :: ants) (_ :: ps) = check ants ps
-      | check (_ :: ants) ps = check ants ps;
-  in check antiqs [] end;
-
-
-(* scan_antiquote *)
-
-open BasicSymbolPos;
-structure T = OuterLex;
-
-local
-
-val scan_txt =
-  $$$ "@" --| Scan.ahead (~$$$ "{") ||
-  Scan.one (fn (s, _) => s <> "@" andalso s <> "\\<lbrace>" andalso s <> "\\<rbrace>"
-    andalso Symbol.is_regular s) >> single;
-
-val scan_ant =
-  T.scan_quoted ||
-  Scan.one (fn (s, _) => s <> "}" andalso Symbol.is_regular s) >> single;
-
-val scan_antiq =
-  SymbolPos.scan_pos -- ($$$ "@" |-- $$$ "{" |--
-    T.!!! "missing closing brace of antiquotation"
-      (Scan.repeat scan_ant -- ($$$ "}" |-- SymbolPos.scan_pos)))
-  >> (fn (pos1, (body, pos2)) => Antiq (flat body, Position.encode_range (pos1, pos2)));
-
-in
-
-val scan_antiquote =
- (Scan.repeat1 scan_txt >> (Text o SymbolPos.content o flat) ||
-  scan_antiq ||
-  SymbolPos.scan_pos --| $$$ "\\<lbrace>" >> Open ||
-  SymbolPos.scan_pos --| $$$ "\\<rbrace>" >> Close);
-
-end;
-
-
-(* read *)
-
-fun report_antiq (Antiq (_, pos)) = Position.report Markup.antiq pos
-  | report_antiq _ = ();
-
-fun read ([], _) = []
-  | read (syms, pos) =
-      (case Scan.read SymbolPos.stopper (Scan.repeat scan_antiquote) syms of
-        SOME xs => (List.app report_antiq xs; check_nesting xs; xs)
-      | NONE => error ("Malformed quotation/antiquotation source" ^ Position.str_of pos));
-
-
-(* read_antiq *)
-
-fun read_antiq lex scan (syms, pos) =
-  let
-    fun err msg = cat_error msg ("Malformed antiquotation" ^ Position.str_of pos ^ ":\n" ^
-      "@{" ^ SymbolPos.content syms ^ "}");
-
-    val res =
-      Source.of_list syms
-      |> T.source' {do_recover = NONE} (K (lex, Scan.empty_lexicon))
-      |> T.source_proper
-      |> Source.source T.stopper (Scan.error (Scan.bulk scan)) NONE
-      |> Source.exhaust;
-  in (case res of [x] => x | _ => err "") handle ERROR msg => err msg end;
-
-end;
--- a/src/Pure/Isar/args.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/args.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -32,7 +32,7 @@
   val mode: string -> bool context_parser
   val maybe: 'a parser -> 'a option parser
   val name_source: string parser
-  val name_source_position: (SymbolPos.text * Position.T) parser
+  val name_source_position: (Symbol_Pos.text * Position.T) parser
   val name: string parser
   val binding: binding parser
   val alt_name: string parser
--- a/src/Pure/Isar/attrib.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/attrib.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -26,7 +26,8 @@
   val add_attributes: (bstring * (src -> attribute) * string) list -> theory -> theory
   val syntax: attribute context_parser -> src -> attribute
   val setup: Binding.binding -> attribute context_parser -> string -> theory -> theory
-  val attribute_setup: bstring * Position.T -> string * Position.T -> string -> theory -> theory
+  val attribute_setup: bstring * Position.T -> Symbol_Pos.text * Position.T -> string ->
+    theory -> theory
   val no_args: attribute -> src -> attribute
   val add_del: attribute -> attribute -> attribute context_parser
   val add_del_args: attribute -> attribute -> src -> attribute
--- a/src/Pure/Isar/class.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/class.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -236,7 +236,7 @@
         thy
         |> Sign.declare_const [] ((b, ty0), syn)
         |> snd
-        |> pair ((Binding.name_of b, ty), (c, ty'))
+        |> pair ((Name.of_binding b, ty), (c, ty'))
       end;
   in
     thy
--- a/src/Pure/Isar/code.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/code.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -98,14 +98,12 @@
     then error ("Code attribute " ^ name ^ " already declared") else add_parser attr attrs)
   end;
 
-val _ =
-  let
-    val code_attr = Scan.peek (fn context =>
-      List.foldr op || Scan.fail (map snd (CodeAttr.get (Context.theory_of context))));
-  in
-    Context.>> (Context.map_theory
-      (Attrib.setup (Binding.name "code") code_attr "declare theorems for code generation"))
-  end;
+val _ = Context.>> (Context.map_theory
+  (Attrib.setup (Binding.name "code")
+    (Scan.peek (fn context =>
+      List.foldr op || Scan.fail (map snd (CodeAttr.get (Context.theory_of context)))))
+    "declare theorems for code generation"));
+
 
 
 (** logical and syntactical specification of executable code **)
--- a/src/Pure/Isar/code_unit.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/code_unit.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -215,10 +215,11 @@
           |> Conjunction.elim_balanced (length thms)
   in
     thms
-    |> burrow_thms (canonical_tvars thy purify_tvar)
     |> map (canonical_vars thy purify_var)
     |> map (canonical_absvars purify_var)
     |> map Drule.zero_var_indexes
+    |> burrow_thms (canonical_tvars thy purify_tvar)
+    |> Drule.zero_var_indexes_list
   end;
 
 
--- a/src/Pure/Isar/constdefs.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/constdefs.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -36,7 +36,7 @@
     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 Binding.name_of d of
+      (case Option.map Name.of_binding d of
         NONE => ()
       | SOME c' =>
           if c <> c' then
--- a/src/Pure/Isar/element.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/element.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -96,7 +96,7 @@
 fun map_ctxt {binding, typ, term, pattern, fact, attrib} =
   fn Fixes fixes => Fixes (fixes |> map (fn (x, T, mx) => (binding x, Option.map typ T, mx)))
    | Constrains xs => Constrains (xs |> map (fn (x, T) =>
-      (Binding.name_of (binding (Binding.name x)), typ T)))
+      (Name.of_binding (binding (Binding.name x)), typ T)))
    | Assumes asms => Assumes (asms |> map (fn ((a, atts), propps) =>
       ((binding a, map attrib atts), propps |> map (fn (t, ps) => (term t, map pattern ps)))))
    | Defines defs => Defines (defs |> map (fn ((a, atts), (t, ps)) =>
--- a/src/Pure/Isar/expression.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/expression.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -125,8 +125,8 @@
 
     val (implicit, expr') = params_expr expr;
 
-    val implicit' = map (#1 #> Binding.name_of) implicit;
-    val fixed' = map (#1 #> Binding.name_of) fixed;
+    val implicit' = map (#1 #> Name.of_binding) implicit;
+    val fixed' = map (#1 #> Name.of_binding) fixed;
     val _ = reject_dups "Duplicate fixed parameter(s): " fixed';
     val implicit'' =
       if strict then []
@@ -203,7 +203,7 @@
 
 fun parse_concl prep_term ctxt concl =
   (map o map) (fn (t, ps) =>
-    (prep_term (ProofContext.set_mode ProofContext.mode_schematic ctxt) t, (* FIXME ?? *)
+    (prep_term (ProofContext.set_mode ProofContext.mode_schematic ctxt) t,
       map (prep_term (ProofContext.set_mode ProofContext.mode_pattern ctxt)) ps)) concl;
 
 
@@ -288,13 +288,13 @@
 fun closeup _ _ false elem = elem
   | closeup ctxt parms true elem =
       let
+        (* FIXME consider closing in syntactic phase -- before type checking *)
         fun close_frees t =
           let
             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; (* FIXME use fold Logic.all *)
-  (* FIXME consider closing in syntactic phase *)
+          in fold (Logic.all o Free) rev_frees t end;
 
         fun no_binds [] = []
           | no_binds _ = error "Illegal term bindings in context element";
@@ -325,9 +325,7 @@
   in (loc, morph) end;
 
 fun finish_elem ctxt parms do_close elem =
-  let
-    val elem' = finish_primitive parms (closeup ctxt parms do_close) elem;
-  in elem' end
+  finish_primitive parms (closeup ctxt parms do_close) elem;
 
 fun finish ctxt parms do_close insts elems =
   let
@@ -352,7 +350,7 @@
     fun prep_insts (loc, (prfx, inst)) (i, insts, ctxt) =
       let
         val (parm_names, parm_types) = Locale.params_of thy loc |>
-          map_split (fn (b, SOME T, _) => (Binding.name_of b, T))
+          map_split (fn (b, SOME T, _) => (Name.of_binding b, T))
             (*FIXME return value of Locale.params_of has strange type*)
         val inst' = prep_inst ctxt parm_names inst;
         val parm_types' = map (TypeInfer.paramify_vars o
@@ -363,7 +361,7 @@
         val inst''' = insts'' |> List.last |> snd |> snd;
         val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst''') ctxt;
         val ctxt'' = Locale.activate_declarations thy (loc, morph) ctxt;
-      in (i+1, insts', ctxt'') end;
+      in (i + 1, insts', ctxt'') end;
 
     fun prep_elem insts raw_elem (elems, ctxt) =
       let
@@ -386,7 +384,7 @@
       prep_concl raw_concl (insts', elems, ctxt5);
 
     (* Retrieve parameter types *)
-    val xs = fold (fn Fixes fixes => (fn ps => ps @ map (Binding.name_of o #1) fixes)
+    val xs = fold (fn Fixes fixes => (fn ps => ps @ map (Name.of_binding o #1) fixes)
       | _ => fn ps => ps) (Fixes fors :: elems') [];
     val (Ts, ctxt7) = fold_map ProofContext.inferred_param xs ctxt6;
     val parms = xs ~~ Ts;  (* params from expression and elements *)
@@ -564,9 +562,7 @@
   in text' end;
 
 fun eval_elem ctxt elem text =
-  let
-    val text' = eval_text ctxt true elem text;
-  in text' end;
+  eval_text ctxt true elem text;
 
 fun eval ctxt deps elems =
   let
@@ -676,7 +672,7 @@
             thy'
             |> Sign.mandatory_path (Binding.name_of aname)
             |> PureThy.note_thmss Thm.internalK
-              [((Binding.name introN, []), [([intro], [Locale.unfold_attrib])])]
+              [((Binding.name introN, []), [([intro], [Locale.unfold_add])])]
             ||> Sign.restore_naming thy';
           in (SOME statement, SOME intro, axioms, thy'') end;
     val (b_pred, b_intro, b_axioms, thy'''') =
@@ -690,7 +686,7 @@
             thy'''
             |> Sign.mandatory_path (Binding.name_of pname)
             |> PureThy.note_thmss Thm.internalK
-                 [((Binding.name introN, []), [([intro], [Locale.intro_attrib])]),
+                 [((Binding.name introN, []), [([intro], [Locale.intro_add])]),
                   ((Binding.name axiomsN, []),
                     [(map (Drule.standard o Element.conclude_witness) axioms, [])])]
             ||> Sign.restore_naming thy''';
@@ -712,7 +708,7 @@
 fun defines_to_notes thy (Defines defs) =
       Notes (Thm.definitionK, map (fn (a, (def, _)) =>
         (a, [([Assumption.assume (cterm_of thy def)],
-          [(Attrib.internal o K) Locale.witness_attrib])])) defs)
+          [(Attrib.internal o K) Locale.witness_add])])) defs)
   | defines_to_notes _ e = e;
 
 fun gen_add_locale prep_decl
@@ -745,11 +741,11 @@
     val asm = if is_some b_statement then b_statement else a_statement;
 
     val notes =
-        if is_some asm
-        then [(Thm.internalK, [((Binding.suffix_name ("_" ^ axiomsN) bname, []),
-          [([Assumption.assume (cterm_of thy' (the asm))],
-            [(Attrib.internal o K) Locale.witness_attrib])])])]
-        else [];
+      if is_some asm
+      then [(Thm.internalK, [((Binding.suffix_name ("_" ^ axiomsN) bname, []),
+        [([Assumption.assume (cterm_of thy' (the asm))],
+          [(Attrib.internal o K) Locale.witness_add])])])]
+      else [];
 
     val notes' = body_elems |>
       map (defines_to_notes thy') |>
@@ -764,8 +760,7 @@
 
     val loc_ctxt = thy'
       |> Locale.register_locale bname (extraTs, params)
-          (asm, rev defs) (a_intro, b_intro) axioms ([], [])
-          (map (fn n => (n, stamp ())) notes |> rev) (map (fn d => (d, stamp ())) deps' |> rev)
+          (asm, rev defs) (a_intro, b_intro) axioms ([], []) (rev notes) (rev deps')
       |> TheoryTarget.init (SOME name)
       |> fold (fn (kind, facts) => LocalTheory.notes kind facts #> snd) notes';
 
@@ -792,11 +787,11 @@
 
     val ((propss, deps, export), goal_ctxt) = prep_expr expression target_ctxt;
 
-    fun after_qed witss = ProofContext.theory (
-      fold2 (fn (name, morph) => fn wits => Locale.add_dependency target
+    fun after_qed witss = ProofContext.theory
+      (fold2 (fn (name, morph) => fn wits => Locale.add_dependency target
         (name, morph $> Element.satisfy_morphism wits $> export)) deps witss #>
       (fn thy => fold_rev Locale.activate_global_facts
-          (Locale.get_registrations thy) thy));
+        (Locale.get_registrations thy) thy));
   in Element.witness_proof after_qed propss goal_ctxt end;
 
 in
--- a/src/Pure/Isar/isar_cmd.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/isar_cmd.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -6,18 +6,18 @@
 
 signature ISAR_CMD =
 sig
-  val global_setup: string * Position.T -> theory -> theory
-  val local_setup: string * Position.T -> Proof.context -> Proof.context
-  val parse_ast_translation: bool * (string * Position.T) -> theory -> theory
-  val parse_translation: bool * (string * Position.T) -> theory -> theory
-  val print_translation: bool * (string * Position.T) -> theory -> theory
-  val typed_print_translation: bool * (string * Position.T) -> theory -> theory
-  val print_ast_translation: bool * (string * Position.T) -> theory -> theory
-  val oracle: bstring * Position.T -> SymbolPos.text * Position.T -> theory -> theory
+  val global_setup: Symbol_Pos.text * Position.T -> theory -> theory
+  val local_setup: Symbol_Pos.text * Position.T -> Proof.context -> Proof.context
+  val parse_ast_translation: bool * (Symbol_Pos.text * Position.T) -> theory -> theory
+  val parse_translation: bool * (Symbol_Pos.text * Position.T) -> theory -> theory
+  val print_translation: bool * (Symbol_Pos.text * Position.T) -> theory -> theory
+  val typed_print_translation: bool * (Symbol_Pos.text * Position.T) -> theory -> theory
+  val print_ast_translation: bool * (Symbol_Pos.text * Position.T) -> theory -> theory
+  val oracle: bstring * Position.T -> Symbol_Pos.text * Position.T -> theory -> theory
   val add_axioms: ((binding * string) * Attrib.src list) list -> theory -> theory
   val add_defs: (bool * bool) * ((binding * 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 ->
+  val declaration: Symbol_Pos.text * Position.T -> local_theory -> local_theory
+  val simproc_setup: string -> string list -> Symbol_Pos.text * Position.T -> string list ->
     local_theory -> local_theory
   val hide_names: bool -> string * xstring list -> theory -> theory
   val have: (Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state
@@ -38,7 +38,7 @@
   val pr: string list * (int option * int option) -> Toplevel.transition -> Toplevel.transition
   val disable_pr: Toplevel.transition -> Toplevel.transition
   val enable_pr: Toplevel.transition -> Toplevel.transition
-  val ml_diag: bool -> string * Position.T -> Toplevel.transition -> Toplevel.transition
+  val ml_diag: bool -> Symbol_Pos.text * Position.T -> Toplevel.transition -> Toplevel.transition
   val cd: Path.T -> Toplevel.transition -> Toplevel.transition
   val pwd: Toplevel.transition -> Toplevel.transition
   val display_drafts: Path.T list -> Toplevel.transition -> Toplevel.transition
@@ -75,10 +75,10 @@
   val print_prop: (string list * string) -> Toplevel.transition -> Toplevel.transition
   val print_term: (string list * string) -> Toplevel.transition -> Toplevel.transition
   val print_type: (string list * string) -> Toplevel.transition -> Toplevel.transition
-  val header_markup: SymbolPos.text * Position.T -> Toplevel.transition -> Toplevel.transition
-  val local_theory_markup: xstring option * (SymbolPos.text * Position.T) ->
+  val header_markup: Symbol_Pos.text * Position.T -> Toplevel.transition -> Toplevel.transition
+  val local_theory_markup: xstring option * (Symbol_Pos.text * Position.T) ->
     Toplevel.transition -> Toplevel.transition
-  val proof_markup: SymbolPos.text * Position.T -> Toplevel.transition -> Toplevel.transition
+  val proof_markup: Symbol_Pos.text * Position.T -> Toplevel.transition -> Toplevel.transition
 end;
 
 structure IsarCmd: ISAR_CMD =
@@ -152,7 +152,7 @@
 
 fun oracle (name, pos) (body_txt, body_pos) =
   let
-    val body = SymbolPos.content (SymbolPos.explode (body_txt, body_pos));
+    val body = Symbol_Pos.content (Symbol_Pos.explode (body_txt, body_pos));
     val txt =
       "local\n\
       \  val binding = " ^ ML_Syntax.make_binding (name, pos) ^ ";\n\
--- a/src/Pure/Isar/isar_syn.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/isar_syn.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -272,7 +272,7 @@
   OuterSyntax.local_theory "lemmas" "define lemmas" K.thy_decl (theorems Thm.lemmaK);
 
 val _ =
-  OuterSyntax.local_theory "declare" "declare theorems (improper)" K.thy_decl
+  OuterSyntax.local_theory "declare" "declare theorems" K.thy_decl
     (P.and_list1 SpecParse.xthms1
       >> (fn args => #2 o Specification.theorems_cmd "" [(Attrib.empty_binding, flat args)]));
 
@@ -295,28 +295,35 @@
 
 (* use ML text *)
 
+fun propagate_env (context as Context.Proof lthy) =
+      Context.Proof (LocalTheory.map_contexts (ML_Context.inherit_env context) lthy)
+  | propagate_env context = context;
+
+fun propagate_env_prf prf = Proof.map_contexts
+  (Context.proof_map (ML_Context.inherit_env (Context.Proof (Proof.context_of prf)))) prf;
+
 val _ =
-  OuterSyntax.command "use" "eval ML text from file" (K.tag_ml K.thy_decl)
-    (P.path >> (Toplevel.generic_theory o ThyInfo.exec_file false));
+  OuterSyntax.command "use" "ML text from file" (K.tag_ml K.thy_decl)
+    (P.path >> (fn path => Toplevel.generic_theory (ThyInfo.exec_file false path #> propagate_env)));
 
 val _ =
-  OuterSyntax.command "ML" "eval ML text within theory" (K.tag_ml K.thy_decl)
+  OuterSyntax.command "ML" "ML text within theory or local theory" (K.tag_ml K.thy_decl)
     (P.ML_source >> (fn (txt, pos) =>
-      Toplevel.generic_theory (ML_Context.exec (fn () => ML_Context.eval true pos txt))));
+      Toplevel.generic_theory
+        (ML_Context.exec (fn () => ML_Context.eval true pos txt) #> propagate_env)));
 
 val _ =
-  OuterSyntax.command "ML_prf" "eval ML text within proof" (K.tag_proof K.prf_decl)
+  OuterSyntax.command "ML_prf" "ML text within proof" (K.tag_proof K.prf_decl)
     (P.ML_source >> (fn (txt, pos) =>
       Toplevel.proof (Proof.map_context (Context.proof_map
-        (ML_Context.exec (fn () => ML_Context.eval true pos txt))) #>
-          (fn prf => Proof.map_contexts (ML_Context.inherit_env (Proof.context_of prf)) prf))));
+        (ML_Context.exec (fn () => ML_Context.eval true pos txt))) #> propagate_env_prf)));
 
 val _ =
-  OuterSyntax.command "ML_val" "eval ML text (diagnostic)" (K.tag_ml K.diag)
+  OuterSyntax.command "ML_val" "diagnostic ML text" (K.tag_ml K.diag)
     (P.ML_source >> IsarCmd.ml_diag true);
 
 val _ =
-  OuterSyntax.command "ML_command" "silently eval ML text (diagnostic)" (K.tag_ml K.diag)
+  OuterSyntax.command "ML_command" "diagnostic ML text (silent)" (K.tag_ml K.diag)
     (P.ML_source >> (Toplevel.no_timing oo IsarCmd.ml_diag false));
 
 val _ =
@@ -397,7 +404,7 @@
 (* locales *)
 
 val locale_val =
-  SpecParse.locale_expression --
+  SpecParse.locale_expression false --
     Scan.optional (P.$$$ "+" |-- P.!!! (Scan.repeat1 SpecParse.context_element)) [] ||
   Scan.repeat1 SpecParse.context_element >> pair ([], []);
 
@@ -411,25 +418,24 @@
 val _ =
   OuterSyntax.command "sublocale"
     "prove sublocale relation between a locale and a locale expression" K.thy_goal
-    (P.xname --| (P.$$$ "\\<subseteq>" || P.$$$ "<") -- P.!!! SpecParse.locale_expression
+    (P.xname --| (P.$$$ "\\<subseteq>" || P.$$$ "<") -- P.!!! (SpecParse.locale_expression false)
       >> (fn (loc, expr) =>
-        Toplevel.print o (Toplevel.theory_to_proof (Expression.sublocale_cmd loc expr))));
+          Toplevel.print o Toplevel.theory_to_proof (Expression.sublocale_cmd loc expr)));
 
 val _ =
   OuterSyntax.command "interpretation"
     "prove interpretation of locale expression in theory" K.thy_goal
-    (P.!!! SpecParse.locale_expression --
-      Scan.optional (P.$$$ "where" |-- P.and_list1 (SpecParse.opt_thm_name ":" -- P.prop)) []
-        >> (fn (expr, equations) => Toplevel.print o
-            Toplevel.theory_to_proof (Expression.interpretation_cmd expr equations)));
+    (P.!!! (SpecParse.locale_expression true) --
+      Scan.optional (P.where_ |-- P.and_list1 (SpecParse.opt_thm_name ":" -- P.prop)) []
+      >> (fn (expr, equations) => Toplevel.print o
+          Toplevel.theory_to_proof (Expression.interpretation_cmd expr equations)));
 
 val _ =
   OuterSyntax.command "interpret"
     "prove interpretation of locale expression in proof context"
     (K.tag_proof K.prf_goal)
-    (P.!!! SpecParse.locale_expression
-        >> (fn expr => Toplevel.print o
-            Toplevel.proof' (fn int => Expression.interpret_cmd expr int)));
+    (P.!!! (SpecParse.locale_expression true)
+      >> (fn expr => Toplevel.print o Toplevel.proof' (Expression.interpret_cmd expr)));
 
 
 (* classes *)
--- a/src/Pure/Isar/local_defs.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/local_defs.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -90,7 +90,7 @@
   let
     val ((bvars, mxs), specs) = defs |> split_list |>> split_list;
     val ((bfacts, atts), rhss) = specs |> split_list |>> split_list;
-    val xs = map Binding.name_of bvars;
+    val xs = map Name.of_binding bvars;
     val names = map2 Thm.def_binding_optional bvars bfacts;
     val eqs = mk_def ctxt (xs ~~ rhss);
     val lhss = map (fst o Logic.dest_equals) eqs;
--- a/src/Pure/Isar/local_theory.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/local_theory.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -23,6 +23,7 @@
   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 map_contexts: (Context.generic -> Context.generic) -> local_theory -> local_theory
   val affirm: local_theory -> local_theory
   val pretty: local_theory -> Pretty.T list
   val abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory ->
@@ -163,6 +164,11 @@
 
 fun target f = #2 o target_result (f #> pair ());
 
+fun map_contexts f =
+  theory (Context.theory_map f) #>
+  target (Context.proof_map f) #>
+  Context.proof_map f;
+
 
 (* basic operations *)
 
--- a/src/Pure/Isar/locale.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/locale.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -1,8 +1,7 @@
 (*  Title:      Pure/Isar/locale.ML
     Author:     Clemens Ballarin, TU Muenchen
 
-Locales -- Isar proof contexts as meta-level predicates, with local
-syntax and implicit structures.
+Locales -- managed Isar proof contexts, based on Pure predicates.
 
 Draws basic ideas from Florian Kammueller's original version of
 locales, but uses the richer infrastructure of Isar instead of the raw
@@ -34,9 +33,9 @@
     (string * sort) list * (binding * typ option * mixfix) list ->
     term option * term list ->
     thm option * thm option -> thm list ->
-    (declaration * stamp) list * (declaration * stamp) list ->
-    ((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * stamp) list ->
-    ((string * morphism) * stamp) list -> theory -> theory
+    declaration list * declaration list ->
+    (string * (Attrib.binding * (thm list * Attrib.src list) list) list) list ->
+    (string * morphism) list -> theory -> theory
   val intern: theory -> xstring -> string
   val extern: theory -> string -> xstring
   val defined: theory -> string -> bool
@@ -64,16 +63,17 @@
   val init: string -> theory -> Proof.context
 
   (* Reasoning about locales *)
-  val witness_attrib: attribute
-  val intro_attrib: attribute
-  val unfold_attrib: attribute
+  val get_witnesses: Proof.context -> thm list
+  val get_intros: Proof.context -> thm list
+  val get_unfolds: Proof.context -> thm list
+  val witness_add: attribute
+  val intro_add: attribute
+  val unfold_add: attribute
   val intro_locales_tac: bool -> Proof.context -> thm list -> tactic
 
   (* Registrations *)
-  val add_registration: string * (morphism * morphism) ->
-    theory -> theory
-  val amend_registration: morphism -> string * morphism ->
-    theory -> theory
+  val add_registration: string * (morphism * morphism) -> theory -> theory
+  val amend_registration: morphism -> string * morphism -> theory -> theory
   val get_registrations: theory -> (string * morphism) list
 
   (* Diagnostic *)
@@ -81,27 +81,6 @@
   val print_locale: theory -> bool -> xstring -> unit
 end;
 
-
-(*** Theorem list extensible via attribute --- to control intro_locales_tac ***)
-
-(* FIXME: consider moving named_thms.ML up in the build hierarchy and using that *)
-functor ThmsFun() =
-struct
-
-structure Data = GenericDataFun
-(
-  type T = thm list;
-  val empty = [];
-  val extend = I;
-  fun merge _ = Thm.merge_thms;
-);
-
-val get = Data.get o Context.Proof;
-val add = Thm.declaration_attribute (Data.map o Thm.add_thm);
-
-end;
-
-
 structure Locale: LOCALE =
 struct
 
@@ -140,7 +119,7 @@
           merge (eq_snd op =) (notes, notes')),
             merge (eq_snd op =) (dependencies, dependencies')));
 
-structure LocalesData = TheoryDataFun
+structure Locales = TheoryDataFun
 (
   type T = locale NameSpace.table;
   val empty = NameSpace.empty_table;
@@ -149,26 +128,29 @@
   fun merge _ = NameSpace.join_tables (K merge_locale);
 );
 
-val intern = NameSpace.intern o #1 o LocalesData.get;
-val extern = NameSpace.extern o #1 o LocalesData.get;
-
-fun get_locale thy = Symtab.lookup (#2 (LocalesData.get thy));
+val intern = NameSpace.intern o #1 o Locales.get;
+val extern = NameSpace.extern o #1 o Locales.get;
 
-fun defined thy = is_some o get_locale thy;
+val get_locale = Symtab.lookup o #2 o Locales.get;
+val defined = Symtab.defined o #2 o Locales.get;
 
-fun the_locale thy name = case get_locale thy name
- of SOME (Loc loc) => loc
-  | NONE => error ("Unknown locale " ^ quote name);
+fun the_locale thy name =
+  (case get_locale thy name of
+    SOME (Loc loc) => loc
+  | NONE => error ("Unknown locale " ^ quote name));
 
 fun register_locale binding parameters spec intros axioms decls notes dependencies thy =
-  thy |> LocalesData.map (NameSpace.define (Sign.naming_of thy) (binding,
-    mk_locale ((parameters, spec, intros, axioms), ((decls, notes), dependencies))) #> snd);
+  thy |> Locales.map (NameSpace.define (Sign.naming_of thy)
+    (binding,
+      mk_locale ((parameters, spec, intros, axioms),
+        ((pairself (map (fn decl => (decl, stamp ()))) decls, map (fn n => (n, stamp ())) notes),
+          map (fn d => (d, stamp ())) dependencies))) #> snd);
 
 fun change_locale name =
-  LocalesData.map o apsnd o Symtab.map_entry name o map_locale o apsnd;
+  Locales.map o apsnd o Symtab.map_entry name o map_locale o apsnd;
 
 fun print_locales thy =
-  Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (LocalesData.get thy)))
+  Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (Locales.get thy)))
   |> Pretty.writeln;
 
 
@@ -181,12 +163,12 @@
 fun axioms_of thy = #axioms o the_locale thy;
 
 fun instance_of thy name morph = params_of thy name |>
-  map ((fn (b, T, _) => Free (Binding.name_of b, the T)) #> Morphism.term morph);
+  map (fn (b, T, _) => Morphism.term morph (Free (Name.of_binding b, the T)));
 
 fun specification_of thy = #spec o the_locale thy;
 
 fun declarations_of thy name = the_locale thy name |>
-  #decls |> apfst (map fst) |> apsnd (map fst);
+  #decls |> pairself (map fst);
 
 fun dependencies_of thy name = the_locale thy name |>
   #dependencies |> map fst;
@@ -206,7 +188,7 @@
 
 datatype 'a delayed = Ready of 'a | ToDo of ('a delayed * 'a delayed);
 
-structure IdentifiersData = GenericDataFun
+structure Identifiers = GenericDataFun
 (
   type T = identifiers delayed;
   val empty = Ready empty;
@@ -220,18 +202,17 @@
   | 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))
-  )));
+  (case Identifiers.get (Context.Theory thy) of
+    Ready _ => NONE
+  | ids => SOME (Context.theory_map (Identifiers.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;
+  let val (Ready ids) = (Identifiers.get o Context.Theory) thy in ids end;
+val put_global_idents = Context.theory_map o Identifiers.put o Ready;
 
 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;
+  let val (Ready ids) = (Identifiers.get o Context.Proof) ctxt in ids end;
+val put_local_idents = Context.proof_map o Identifiers.put o Ready;
 
 end;
 
@@ -385,7 +366,7 @@
 
 (*** Registrations: interpretations in theories ***)
 
-structure RegistrationsData = TheoryDataFun
+structure Registrations = TheoryDataFun
 (
   type T = ((string * (morphism * morphism)) * stamp) list;
     (* FIXME mixins need to be stamped *)
@@ -398,17 +379,17 @@
 );
 
 val get_registrations =
-  RegistrationsData.get #> map fst #> map (apsnd op $>);
+  Registrations.get #> map fst #> map (apsnd op $>);
 
 fun add_registration (name, (base_morph, export)) thy =
   roundup thy (fn _ => fn (name', morph') =>
-    (RegistrationsData.map o cons) ((name', (morph', export)), stamp ()))
+    (Registrations.map o cons) ((name', (morph', export)), stamp ()))
     (name, base_morph) (get_global_idents thy, thy) |> snd
     (* FIXME |-> put_global_idents ?*);
 
 fun amend_registration morph (name, base_morph) thy =
   let
-    val regs = (RegistrationsData.get #> map fst) thy;
+    val regs = (Registrations.get #> map fst) thy;
     val base = instance_of thy name base_morph;
     fun match (name', (morph', _)) =
       name = name' andalso eq_list (op aconv) (base, instance_of thy name' morph');
@@ -418,7 +399,7 @@
         space_implode " " (map (quote o Syntax.string_of_term_global thy) base) ^ " available.")
       else ();
   in
-    RegistrationsData.map (nth_map (length regs - 1 - i)
+    Registrations.map (nth_map (length regs - 1 - i)
       (fn ((name, (base, export)), stamp) => ((name, (base $> morph, export)), stamp))) thy
   end;
 
@@ -463,6 +444,7 @@
 
 end;
 
+
 (* Dependencies *)
 
 fun add_dependency loc dep = (change_locale loc o apsnd) (cons (dep, stamp ()));
@@ -470,21 +452,36 @@
 
 (*** Reasoning about locales ***)
 
-(** Storage for witnesses, intro and unfold rules **)
+(* Storage for witnesses, intro and unfold rules *)
 
-structure Witnesses = ThmsFun();
-structure Intros = ThmsFun();
-structure Unfolds = ThmsFun();
+structure Thms = GenericDataFun
+(
+  type T = thm list * thm list * thm list;
+  val empty = ([], [], []);
+  val extend = I;
+  fun merge _ ((witnesses1, intros1, unfolds1), (witnesses2, intros2, unfolds2)) =
+   (Thm.merge_thms (witnesses1, witnesses2),
+    Thm.merge_thms (intros1, intros2),
+    Thm.merge_thms (unfolds1, unfolds2));
+);
 
-val witness_attrib = Witnesses.add;
-val intro_attrib = Intros.add;
-val unfold_attrib = Unfolds.add;
+val get_witnesses = #1 o Thms.get o Context.Proof;
+val get_intros = #2 o Thms.get o Context.Proof;
+val get_unfolds = #3 o Thms.get o Context.Proof;
 
-(** Tactic **)
+val witness_add =
+  Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (Thm.add_thm th x, y, z)));
+val intro_add =
+  Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, Thm.add_thm th y, z)));
+val unfold_add =
+  Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, y, Thm.add_thm th z)));
 
-fun intro_locales_tac eager ctxt facts st =
+
+(* Tactic *)
+
+fun intro_locales_tac eager ctxt =
   Method.intros_tac
-    (Witnesses.get ctxt @ Intros.get ctxt @ (if eager then Unfolds.get ctxt else [])) facts st;
+    (get_witnesses ctxt @ get_intros ctxt @ (if eager then get_unfolds ctxt else []));
 
 val _ = Context.>> (Context.map_theory
  (Method.setup (Binding.name "intro_locales") (Scan.succeed (METHOD o intro_locales_tac false))
--- a/src/Pure/Isar/method.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/method.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -81,7 +81,8 @@
     -> theory -> theory
   val syntax: 'a context_parser -> src -> Proof.context -> 'a * Proof.context
   val setup: binding -> (Proof.context -> method) context_parser -> string -> theory -> theory
-  val method_setup: bstring * Position.T -> string * Position.T -> string -> theory -> theory
+  val method_setup: bstring * Position.T -> Symbol_Pos.text * Position.T -> string ->
+    theory -> theory
   val simple_args: 'a parser -> ('a -> Proof.context -> method) -> src -> Proof.context -> method
   val ctxt_args: (Proof.context -> method) -> src -> Proof.context -> method
   val no_args: method -> src -> Proof.context -> method
--- a/src/Pure/Isar/obtain.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/obtain.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -119,7 +119,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 (Binding.name_of o #1) vars;
+    val xs = map (Name.of_binding o #1) vars;
 
     (*obtain asms*)
     val (asms_ctxt, proppss) = prep_propp (fix_ctxt, map snd raw_asms);
@@ -258,7 +258,7 @@
 
 fun inferred_type (binding, _, mx) ctxt =
   let
-    val x = Binding.name_of binding;
+    val x = Name.of_binding binding;
     val (T, ctxt') = ProofContext.inferred_param x ctxt
   in ((x, T, mx), ctxt') end;
 
--- a/src/Pure/Isar/outer_keyword.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/outer_keyword.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -150,23 +150,25 @@
   Pretty.mark (Markup.command_decl name (kind_of kind))
     (Pretty.str ("Outer syntax command: " ^ quote name ^ " (" ^ kind_of kind ^ ")"));
 
+fun status_writeln s = (Output.status s; writeln s);
+
 fun report () =
   let val (keywords, commands) = CRITICAL (fn () =>
     (dest_keywords (), sort_wrt #1 (Symtab.dest (get_commands ()))))
   in map report_keyword keywords @ map report_command commands end
-  |> Pretty.chunks |> Pretty.string_of |> Output.status;
+  |> Pretty.chunks |> Pretty.string_of |> status_writeln;
 
 
 (* augment tables *)
 
 fun keyword name =
  (change_lexicons (apfst (Scan.extend_lexicon (Symbol.explode name)));
-  Output.status (Pretty.string_of (report_keyword name)));
+  status_writeln (Pretty.string_of (report_keyword name)));
 
 fun command name kind =
  (change_lexicons (apsnd (Scan.extend_lexicon (Symbol.explode name)));
   change_commands (Symtab.update (name, kind));
-  Output.status (Pretty.string_of (report_command (name, kind))));
+  status_writeln (Pretty.string_of (report_command (name, kind))));
 
 
 (* command categories *)
--- a/src/Pure/Isar/outer_lex.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/outer_lex.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -35,7 +35,7 @@
   val is_blank: token -> bool
   val is_newline: token -> bool
   val source_of: token -> string
-  val source_position_of: token -> SymbolPos.text * Position.T
+  val source_position_of: token -> Symbol_Pos.text * Position.T
   val content_of: token -> string
   val unparse: token -> string
   val text_of: token -> string * string
@@ -50,14 +50,15 @@
   val assign: value option -> token -> unit
   val closure: token -> token
   val ident_or_symbolic: string -> bool
-  val !!! : string -> (SymbolPos.T list -> 'a) -> SymbolPos.T list -> 'a
-  val scan_quoted: SymbolPos.T list -> SymbolPos.T list * SymbolPos.T list
+  val !!! : string -> (Symbol_Pos.T list -> 'a) -> Symbol_Pos.T list -> 'a
   val source_proper: (token, 'a) Source.source -> (token, (token, 'a) Source.source) Source.source
   val source': {do_recover: bool Option.option} -> (unit -> Scan.lexicon * Scan.lexicon) ->
-    (SymbolPos.T, 'a) Source.source -> (token, (SymbolPos.T, 'a) Source.source) Source.source
+    (Symbol_Pos.T, 'a) Source.source -> (token, (Symbol_Pos.T, 'a) Source.source) Source.source
   val source: {do_recover: bool Option.option} -> (unit -> Scan.lexicon * Scan.lexicon) ->
     Position.T -> (Symbol.symbol, 'a) Source.source -> (token,
-      (SymbolPos.T, Position.T * (Symbol.symbol, 'a) Source.source) Source.source) Source.source
+      (Symbol_Pos.T, Position.T * (Symbol.symbol, 'a) Source.source) Source.source) Source.source
+  val read_antiq: Scan.lexicon -> (token list -> 'a * token list) ->
+    Symbol_Pos.T list * Position.T -> 'a
 end;
 
 structure OuterLex: OUTER_LEX =
@@ -92,7 +93,7 @@
   Nat | String | AltString | Verbatim | Space | Comment | InternalValue |
   Malformed | Error of string | Sync | EOF;
 
-datatype token = Token of (SymbolPos.text * Position.range) * (token_kind * string) * slot;
+datatype token = Token of (Symbol_Pos.text * Position.range) * (token_kind * string) * slot;
 
 val str_of_kind =
  fn Command => "command"
@@ -259,11 +260,9 @@
 
 (** scanners **)
 
-open BasicSymbolPos;
+open Basic_Symbol_Pos;
 
-fun !!! msg = SymbolPos.!!! ("Outer lexical error: " ^ msg);
-
-fun change_prompt scan = Scan.prompt "# " scan;
+fun !!! msg = Symbol_Pos.!!! ("Outer lexical error: " ^ msg);
 
 
 (* scan symbolic idents *)
@@ -286,36 +285,6 @@
   | ident_or_symbolic s = Syntax.is_identifier s orelse is_symid s;
 
 
-(* scan strings *)
-
-local
-
-val char_code =
-  Scan.one (Symbol.is_ascii_digit o symbol) --
-  Scan.one (Symbol.is_ascii_digit o symbol) --
-  Scan.one (Symbol.is_ascii_digit o symbol) :|--
-  (fn (((a, pos), (b, _)), (c, _)) =>
-    let val (n, _) = Library.read_int [a, b, c]
-    in if n <= 255 then Scan.succeed [(chr n, pos)] else Scan.fail end);
-
-fun scan_str q =
-  $$$ "\\" |-- !!! "bad escape character in string" ($$$ q || $$$ "\\" || char_code) ||
-  Scan.one (fn (s, _) => s <> q andalso s <> "\\" andalso Symbol.is_regular s) >> single;
-
-fun scan_strs q =
-  (SymbolPos.scan_pos --| $$$ q) -- !!! "missing quote at end of string"
-    (change_prompt ((Scan.repeat (scan_str q) >> flat) -- ($$$ q |-- SymbolPos.scan_pos)));
-
-in
-
-val scan_string = scan_strs "\"";
-val scan_alt_string = scan_strs "`";
-
-val scan_quoted = Scan.trace (scan_string || scan_alt_string) >> #2;
-
-end;
-
-
 (* scan verbatim text *)
 
 val scan_verb =
@@ -323,8 +292,9 @@
   Scan.one (fn (s, _) => s <> "*" andalso Symbol.is_regular s) >> single;
 
 val scan_verbatim =
-  (SymbolPos.scan_pos --| $$$ "{" --| $$$ "*") -- !!! "missing end of verbatim text"
-    (change_prompt ((Scan.repeat scan_verb >> flat) -- ($$$ "*" |-- $$$ "}" |-- SymbolPos.scan_pos)));
+  (Symbol_Pos.scan_pos --| $$$ "{" --| $$$ "*") -- !!! "missing end of verbatim text"
+    (Symbol_Pos.change_prompt
+      ((Scan.repeat scan_verb >> flat) -- ($$$ "*" |-- $$$ "}" |-- Symbol_Pos.scan_pos)));
 
 
 (* scan space *)
@@ -339,14 +309,14 @@
 (* scan comment *)
 
 val scan_comment =
-  SymbolPos.scan_pos -- (SymbolPos.scan_comment_body !!! -- SymbolPos.scan_pos);
+  Symbol_Pos.scan_pos -- (Symbol_Pos.scan_comment_body !!! -- Symbol_Pos.scan_pos);
 
 
 (* scan malformed symbols *)
 
 val scan_malformed =
   $$$ Symbol.malformed |--
-    change_prompt (Scan.many (Symbol.is_regular o symbol))
+    Symbol_Pos.change_prompt (Scan.many (Symbol.is_regular o symbol))
   --| Scan.option ($$$ Symbol.end_malformed);
 
 
@@ -360,14 +330,14 @@
 fun token_leq ((_, syms1), (_, syms2)) = length syms1 <= length syms2;
 
 fun token k ss =
-  Token ((SymbolPos.implode ss, SymbolPos.range ss), (k, SymbolPos.untabify_content ss), Slot);
+  Token ((Symbol_Pos.implode ss, Symbol_Pos.range ss), (k, Symbol_Pos.untabify_content ss), Slot);
 
 fun token_range k (pos1, (ss, pos2)) =
-  Token (SymbolPos.implode_range pos1 pos2 ss, (k, SymbolPos.untabify_content ss), Slot);
+  Token (Symbol_Pos.implode_range pos1 pos2 ss, (k, Symbol_Pos.untabify_content ss), Slot);
 
 fun scan (lex1, lex2) = !!! "bad input"
-  (scan_string >> token_range String ||
-    scan_alt_string >> token_range AltString ||
+  (Symbol_Pos.scan_string >> token_range String ||
+    Symbol_Pos.scan_alt_string >> token_range AltString ||
     scan_verbatim >> token_range Verbatim ||
     scan_comment >> token_range Comment ||
     scan_space >> token Space ||
@@ -392,13 +362,29 @@
 in
 
 fun source' {do_recover} get_lex =
-  Source.source SymbolPos.stopper (Scan.bulk (fn xs => scan (get_lex ()) xs))
+  Source.source Symbol_Pos.stopper (Scan.bulk (fn xs => scan (get_lex ()) xs))
     (Option.map (rpair recover) do_recover);
 
 fun source do_recover get_lex pos src =
-  SymbolPos.source pos src
+  Symbol_Pos.source pos src
   |> source' do_recover get_lex;
 
 end;
 
+
+(* read_antiq *)
+
+fun read_antiq lex scan (syms, pos) =
+  let
+    fun err msg = cat_error msg ("Malformed antiquotation" ^ Position.str_of pos ^ ":\n" ^
+      "@{" ^ Symbol_Pos.content syms ^ "}");
+
+    val res =
+      Source.of_list syms
+      |> source' {do_recover = NONE} (K (lex, Scan.empty_lexicon))
+      |> source_proper
+      |> Source.source stopper (Scan.error (Scan.bulk scan)) NONE
+      |> Source.exhaust;
+  in (case res of [x] => x | _ => err "") handle ERROR msg => err msg end;
+
 end;
--- a/src/Pure/Isar/outer_parse.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/outer_parse.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -84,8 +84,8 @@
   val fixes: (binding * string option * mixfix) list parser
   val for_fixes: (binding * string option * mixfix) list parser
   val for_simple_fixes: (binding * string option) list parser
-  val ML_source: (SymbolPos.text * Position.T) parser
-  val doc_source: (SymbolPos.text * Position.T) parser
+  val ML_source: (Symbol_Pos.text * Position.T) parser
+  val doc_source: (Symbol_Pos.text * Position.T) parser
   val term_group: string parser
   val prop_group: string parser
   val term: string parser
--- a/src/Pure/Isar/outer_syntax.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/outer_syntax.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -224,7 +224,7 @@
 type isar =
   (Toplevel.transition, (Toplevel.transition option,
     (OuterLex.token, (OuterLex.token option, (OuterLex.token, (OuterLex.token,
-      (SymbolPos.T, Position.T * (Symbol.symbol, (string, unit) Source.source)
+      (Symbol_Pos.T, Position.T * (Symbol.symbol, (string, unit) Source.source)
   Source.source) Source.source) Source.source) Source.source)
   Source.source) Source.source) Source.source) Source.source;
 
--- a/src/Pure/Isar/proof_context.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/proof_context.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -36,7 +36,9 @@
   val theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context
   val extern_fact: Proof.context -> string -> xstring
   val pretty_term_abbrev: Proof.context -> term -> Pretty.T
-  val pretty_thm_legacy: thm -> Pretty.T
+  val pretty_thm_aux: Proof.context -> bool -> thm -> Pretty.T
+  val pretty_thms_aux: Proof.context -> bool -> thm list -> Pretty.T
+  val pretty_fact_aux: Proof.context -> bool -> string * thm list -> Pretty.T
   val pretty_thm: Proof.context -> thm -> Pretty.T
   val pretty_thms: Proof.context -> thm list -> Pretty.T
   val pretty_fact: Proof.context -> string * thm list -> Pretty.T
@@ -296,25 +298,28 @@
 
 fun pretty_term_abbrev ctxt = Syntax.pretty_term (set_mode mode_abbrev ctxt);
 
-fun pretty_thm_legacy th =
-  let val thy = Thm.theory_of_thm th
-  in Display.pretty_thm_aux (Syntax.pp_global thy) true false [] th end;
+fun pretty_thm_aux ctxt show_status th =
+  let
+    val flags = {quote = false, show_hyps = true, show_status = show_status};
+    val asms = map Thm.term_of (Assumption.all_assms_of ctxt);
+  in Display.pretty_thm_aux (Syntax.pp ctxt) flags asms th end;
 
-fun pretty_thm ctxt th =
-  let val asms = map Thm.term_of (Assumption.all_assms_of ctxt)
-  in Display.pretty_thm_aux (Syntax.pp ctxt) false true asms th end;
-
-fun pretty_thms ctxt [th] = pretty_thm ctxt th
-  | pretty_thms ctxt ths = Pretty.blk (0, Pretty.fbreaks (map (pretty_thm ctxt) ths));
+fun pretty_thms_aux ctxt flag [th] = pretty_thm_aux ctxt flag th
+  | pretty_thms_aux ctxt flag ths =
+      Pretty.blk (0, Pretty.fbreaks (map (pretty_thm_aux ctxt flag) ths));
 
 fun pretty_fact_name ctxt a = Pretty.block
   [Pretty.markup (Markup.fact a) [Pretty.str (extern_fact ctxt a)], Pretty.str ":"];
 
-fun pretty_fact ctxt ("", ths) = pretty_thms ctxt ths
-  | pretty_fact ctxt (a, [th]) = Pretty.block
-      [pretty_fact_name ctxt a, Pretty.brk 1, pretty_thm ctxt th]
-  | pretty_fact ctxt (a, ths) = Pretty.block
-      (Pretty.fbreaks (pretty_fact_name ctxt a :: map (pretty_thm ctxt) ths));
+fun pretty_fact_aux ctxt flag ("", ths) = pretty_thms_aux ctxt flag ths
+  | pretty_fact_aux ctxt flag (a, [th]) = Pretty.block
+      [pretty_fact_name ctxt a, Pretty.brk 1, pretty_thm_aux ctxt flag th]
+  | pretty_fact_aux ctxt flag (a, ths) = Pretty.block
+      (Pretty.fbreaks (pretty_fact_name ctxt a :: map (pretty_thm_aux ctxt flag) ths));
+
+fun pretty_thm ctxt = pretty_thm_aux ctxt true;
+fun pretty_thms ctxt = pretty_thms_aux ctxt true;
+fun pretty_fact ctxt = pretty_fact_aux ctxt true;
 
 val string_of_thm = Pretty.string_of oo pretty_thm;
 
@@ -430,7 +435,7 @@
 
 local
 
-val token_content = Syntax.read_token #>> SymbolPos.content;
+val token_content = Syntax.read_token #>> Symbol_Pos.content;
 
 fun prep_const_proper ctxt (c, pos) =
   let val t as (Const (d, _)) =
@@ -1003,7 +1008,7 @@
 fun prep_vars prep_typ internal =
   fold_map (fn (raw_b, raw_T, raw_mx) => fn ctxt =>
     let
-      val raw_x = Binding.name_of raw_b;
+      val raw_x = Name.of_binding 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);
@@ -1113,7 +1118,7 @@
 fun gen_fixes prep raw_vars ctxt =
   let
     val (vars, _) = prep raw_vars ctxt;
-    val (xs', ctxt') = Variable.add_fixes (map (Binding.name_of o #1) vars) ctxt;
+    val (xs', ctxt') = Variable.add_fixes (map (Name.of_binding o #1) vars) ctxt;
     val ctxt'' =
       ctxt'
       |> fold_map declare_var (map2 (fn x' => fn (_, T, mx) => (x', T, mx)) xs' vars)
--- a/src/Pure/Isar/proof_display.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/proof_display.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -6,12 +6,12 @@
 
 signature PROOF_DISPLAY =
 sig
-  val pprint_context: Proof.context -> pprint_args -> unit
-  val pprint_typ: theory -> typ -> pprint_args -> unit
-  val pprint_term: theory -> term -> pprint_args -> unit
-  val pprint_ctyp: ctyp -> pprint_args -> unit
-  val pprint_cterm: cterm -> pprint_args -> unit
-  val pprint_thm: thm -> pprint_args -> unit
+  val pp_context: Proof.context -> Pretty.T
+  val pp_thm: thm -> Pretty.T
+  val pp_typ: theory -> typ -> Pretty.T
+  val pp_term: theory -> term -> Pretty.T
+  val pp_ctyp: ctyp -> Pretty.T
+  val pp_cterm: cterm -> Pretty.T
   val print_theorems_diff: theory -> theory -> unit
   val print_theorems: theory -> unit
   val pretty_full_theory: bool -> theory -> Pretty.T
@@ -26,18 +26,22 @@
 
 (* toplevel pretty printing *)
 
-fun pprint_context ctxt = Pretty.pprint
+fun pp_context ctxt =
  (if ! ProofContext.debug then
     Pretty.quote (Pretty.big_list "proof context:" (ProofContext.pretty_context ctxt))
   else Pretty.str "<context>");
 
-fun pprint pretty thy = Pretty.pprint o Pretty.quote o pretty (Syntax.init_pretty_global thy);
+fun pp_thm th =
+  let
+    val thy = Thm.theory_of_thm th;
+    val flags = {quote = true, show_hyps = false, show_status = true};
+  in Display.pretty_thm_aux (Syntax.pp_global thy) flags [] th end;
 
-val pprint_typ = pprint Syntax.pretty_typ;
-val pprint_term = pprint Syntax.pretty_term;
-fun pprint_ctyp cT = pprint_typ (Thm.theory_of_ctyp cT) (Thm.typ_of cT);
-fun pprint_cterm ct = pprint_term (Thm.theory_of_cterm ct) (Thm.term_of ct);
-val pprint_thm = Pretty.pprint o ProofContext.pretty_thm_legacy;
+val pp_typ = Pretty.quote oo Syntax.pretty_typ_global;
+val pp_term = Pretty.quote oo Syntax.pretty_term_global;
+
+fun pp_ctyp cT = pp_typ (Thm.theory_of_ctyp cT) (Thm.typ_of cT);
+fun pp_cterm ct = pp_term (Thm.theory_of_cterm ct) (Thm.term_of ct);
 
 
 (* theorems and theory *)
@@ -64,7 +68,7 @@
 
 fun pretty_rule ctxt s thm =
   Pretty.block [Pretty.str (s ^ " attempt to solve goal by exported rule:"),
-    Pretty.fbrk, ProofContext.pretty_thm ctxt thm];
+    Pretty.fbrk, ProofContext.pretty_thm_aux ctxt false thm];
 
 val string_of_rule = Pretty.string_of ooo pretty_rule;
 
@@ -79,7 +83,7 @@
 
 fun pretty_facts ctxt =
   flat o (separate [Pretty.fbrk, Pretty.str "and "]) o
-    map (single o ProofContext.pretty_fact ctxt);
+    map (single o ProofContext.pretty_fact_aux ctxt false);
 
 in
 
@@ -90,7 +94,7 @@
   else Pretty.writeln
     (case facts of
       [fact] => Pretty.blk (1, [pretty_fact_name (kind, name), Pretty.fbrk,
-        ProofContext.pretty_fact ctxt fact])
+        ProofContext.pretty_fact_aux ctxt false fact])
     | _ => Pretty.blk (1, [pretty_fact_name (kind, name), Pretty.fbrk,
         Pretty.blk (1, Pretty.str "(" :: pretty_facts ctxt facts @ [Pretty.str ")"])]));
 
--- a/src/Pure/Isar/rule_insts.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/rule_insts.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -220,8 +220,11 @@
 
 in
 
-fun where_att x = (Scan.lift (P.and_list inst) >> (fn args =>
-  Thm.rule_attribute (fn context => read_instantiate_mixed (Context.proof_of context) args))) x;
+val _ = Context.>> (Context.map_theory
+  (Attrib.setup (Binding.name "where")
+    (Scan.lift (P.and_list inst) >> (fn args =>
+      Thm.rule_attribute (fn context => read_instantiate_mixed (Context.proof_of context) args)))
+    "named instantiation of theorem"));
 
 end;
 
@@ -243,19 +246,15 @@
 
 in
 
-fun of_att x = (Scan.lift insts >> (fn args =>
-  Thm.rule_attribute (fn context => read_instantiate_mixed' (Context.proof_of context) args))) x;
+val _ = Context.>> (Context.map_theory
+  (Attrib.setup (Binding.name "of")
+    (Scan.lift insts >> (fn args =>
+      Thm.rule_attribute (fn context => read_instantiate_mixed' (Context.proof_of context) args)))
+    "positional instantiation of theorem"));
 
 end;
 
 
-(* setup *)
-
-val _ = Context.>> (Context.map_theory
- (Attrib.setup (Binding.name "where") where_att "named instantiation of theorem" #>
-  Attrib.setup (Binding.name "of") of_att "positional instantiation of theorem"));
-
-
 
 (** tactics **)
 
--- a/src/Pure/Isar/spec_parse.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/spec_parse.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -22,7 +22,7 @@
   val locale_fixes: (binding * string option * mixfix) list parser
   val locale_insts: (string option list * (Attrib.binding * string) list) parser
   val class_expr: string list parser
-  val locale_expression: Expression.expression parser
+  val locale_expression: bool -> Expression.expression parser
   val locale_keyword: string parser
   val context_element: Element.context parser
   val statement: (Attrib.binding * (string * string list) list) list parser
@@ -77,7 +77,7 @@
 
 val locale_insts =
   Scan.optional (P.$$$ "[" |-- P.!!! (Scan.repeat1 (P.maybe P.term) --| P.$$$ "]")) []
-  -- Scan.optional (P.$$$ "where" |-- P.and_list1 (opt_thm_name ":" -- P.prop)) [];
+  -- Scan.optional (P.where_ |-- P.and_list1 (opt_thm_name ":" -- P.prop)) [];
 
 local
 
@@ -95,13 +95,12 @@
 fun plus1_unless test scan =
   scan ::: Scan.repeat (P.$$$ "+" |-- Scan.unless test (P.!!! scan));
 
-val rename = P.name -- Scan.option P.mixfix;
+fun prefix mandatory =
+  P.name -- (P.$$$ "!" >> K true || P.$$$ "?" >> K false || Scan.succeed mandatory) --| P.$$$ ":";
 
-val prefix = P.name -- Scan.optional (P.$$$ "!" >> K true) false --| P.$$$ ":";
-val named = P.name -- (P.$$$ "=" |-- P.term);
-val position = P.maybe P.term;
-val instance = P.$$$ "where" |-- P.and_list1 named >> Expression.Named ||
-  Scan.repeat1 position >> Expression.Positional;
+val instance = P.where_ |--
+  P.and_list1 (P.name -- (P.$$$ "=" |-- P.term)) >> Expression.Named ||
+  Scan.repeat1 (P.maybe P.term) >> Expression.Positional;
 
 in
 
@@ -110,12 +109,12 @@
 
 val class_expr = plus1_unless locale_keyword P.xname;
 
-val locale_expression =
+fun locale_expression mandatory =
   let
-    fun expr2 x = P.xname x;
-    fun expr1 x = (Scan.optional prefix ("", false) -- expr2 --
-      Scan.optional instance (Expression.Named []) >> (fn ((p, l), i) => (l, (p, i)))) x;
-    fun expr0 x = (plus1_unless locale_keyword expr1) x;
+    val expr2 = P.xname;
+    val expr1 = Scan.optional (prefix mandatory) ("", false) -- expr2 --
+      Scan.optional instance (Expression.Named []) >> (fn ((p, l), i) => (l, (p, i)));
+    val expr0 = plus1_unless locale_keyword expr1;
   in expr0 -- Scan.optional (P.$$$ "for" |-- P.!!! locale_fixes) [] end;
 
 val context_element = P.group "context element" loc_element;
@@ -128,7 +127,7 @@
 val statement = P.and_list1 (opt_thm_name ":" -- Scan.repeat1 P.propp);
 
 val obtain_case =
-  P.parbinding -- (Scan.optional (P.simple_fixes --| P.$$$ "where") [] --
+  P.parbinding -- (Scan.optional (P.simple_fixes --| P.where_) [] --
     (P.and_list1 (Scan.repeat1 P.prop) >> flat));
 
 val general_statement =
--- a/src/Pure/Isar/specification.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/specification.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -161,7 +161,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), _) => (Binding.name_of b, T)) vars;
+    val xs = map (fn ((b, T), _) => (Name.of_binding b, T)) vars;
 
     (*consts*)
     val (consts, consts_thy) = thy |> fold_map (Theory.specify_const []) vars;
@@ -171,7 +171,7 @@
     val (axioms, axioms_thy) = (specs, consts_thy) |-> fold_map (fn ((b, atts), props) =>
         fold_map Thm.add_axiom
           (map (apfst (fn a => Binding.map_name (K a) b))
-            (PureThy.name_multi (Binding.name_of b) (map subst props)))
+            (PureThy.name_multi (Name.of_binding b) (map subst props)))
         #>> (fn ths => ((b, atts), [(map Drule.standard' ths, [])])));
 
     (*facts*)
@@ -198,7 +198,7 @@
         [] => (Binding.name x, NoSyn)
       | [((b, _), mx)] =>
           let
-            val y = Binding.name_of b;
+            val y = Name.of_binding b;
             val _ = x = y orelse
               error ("Head of definition " ^ quote x ^ " differs from declaration " ^ quote y ^
                 Position.str_of (Binding.pos_of b));
@@ -234,7 +234,7 @@
         [] => (Binding.name x, NoSyn)
       | [((b, _), mx)] =>
           let
-            val y = Binding.name_of b;
+            val y = Name.of_binding b;
             val _ = x = y orelse
               error ("Head of abbreviation " ^ quote x ^ " differs from declaration " ^ quote y ^
                 Position.str_of (Binding.pos_of b));
@@ -292,10 +292,10 @@
   | Element.Obtains obtains =>
       let
         val case_names = obtains |> map_index (fn (i, (b, _)) =>
-          if Binding.is_empty b then string_of_int (i + 1) else Binding.name_of b);
+          if Binding.is_empty b then string_of_int (i + 1) else Name.of_binding b);
         val constraints = obtains |> map (fn (_, (vars, _)) =>
           Element.Constrains
-            (vars |> map_filter (fn (x, SOME T) => SOME (Binding.name_of x, T) | _ => NONE)));
+            (vars |> map_filter (fn (x, SOME T) => SOME (Name.of_binding x, T) | _ => NONE)));
 
         val raw_propp = obtains |> map (fn (_, (_, props)) => map (rpair []) props);
         val (propp, elems_ctxt) = prep_stmt (elems @ constraints) raw_propp ctxt;
@@ -305,7 +305,7 @@
         fun assume_case ((name, (vars, _)), asms) ctxt' =
           let
             val bs = map fst vars;
-            val xs = map Binding.name_of bs;
+            val xs = map Name.of_binding bs;
             val props = map fst asms;
             val (Ts, _) = ctxt'
               |> fold Variable.declare_term props
--- a/src/Pure/Isar/toplevel.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Isar/toplevel.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -311,7 +311,7 @@
 fun controlled_execution f =
   f
   |> debugging
-  |> interruptible;
+  |> Future.interruptible_task;
 
 fun program f =
  (f
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/install_pp_polyml-experimental.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -0,0 +1,18 @@
+(*  Title:      Pure/ML-Systems/install_pp_polyml-experimental.ML
+
+Extra toplevel pretty-printing for Poly/ML; experimental version for
+Poly/ML 5.3.
+*)
+
+addPrettyPrinter (fn depth => fn pretty => fn x =>
+  (case Future.peek x of
+    NONE => PrettyString "<future>"
+  | SOME (Exn.Exn _) => PrettyString "<failed>"
+  | SOME (Exn.Result y) => pretty (y, depth)));
+
+addPrettyPrinter (fn depth => fn pretty => fn x =>
+  (case Lazy.peek x of
+    NONE => PrettyString "<lazy>"
+  | SOME (Exn.Exn _) => PrettyString "<failed>"
+  | SOME (Exn.Result y) => pretty (y, depth)));
+
--- a/src/Pure/ML-Systems/ml_name_space.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML-Systems/ml_name_space.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -4,7 +4,7 @@
 ML name space -- dummy version of Poly/ML 5.2 facility.
 *)
 
-structure ML_NameSpace =
+structure ML_Name_Space =
 struct
 
 type valueVal = unit;
@@ -14,7 +14,7 @@
 type signatureVal = unit;
 type functorVal = unit;
 
-type nameSpace =
+type T =
  {lookupVal:    string -> valueVal option,
   lookupType:   string -> typeVal option,
   lookupFix:    string -> fixityVal option,
@@ -34,7 +34,7 @@
   allSig:       unit -> (string * signatureVal) list,
   allFunct:     unit -> (string * functorVal) list};
 
-val global: nameSpace =
+val global: T =
  {lookupVal = fn _ => NONE,
   lookupType = fn _ => NONE,
   lookupFix = fn _ => NONE,
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/ml_pretty.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -0,0 +1,16 @@
+(*  Title:      Pure/ML-Systems/ml_pretty.ML
+    Author:     Makarius
+
+Raw datatype for ML pretty printing.
+*)
+
+structure ML_Pretty =
+struct
+
+datatype pretty =
+  Block of (string * string) * pretty list * int |
+  String of string * int |
+  Break of bool * int;
+
+end;
+
--- a/src/Pure/ML-Systems/mosml.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML-Systems/mosml.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -45,6 +45,8 @@
 use "ML-Systems/multithreading.ML";
 use "ML-Systems/time_limit.ML";
 use "ML-Systems/ml_name_space.ML";
+use "ML-Systems/ml_pretty.ML";
+use "ML-Systems/use_context.ML";
 
 
 (*low-level pointer equality*)
@@ -118,10 +120,8 @@
     Meta.printLength := n);
 end;
 
-(*interface for toplevel pretty printers, see also Pure/pure_setup.ML*)
-(*the documentation refers to an installPP but I couldn't fine it!*)
-fun make_pp path pprint = ();
-fun install_pp _ = ();
+(*dummy implementation*)
+fun toplevel_pp _ _ _ = ();
 
 (*dummy implementation*)
 fun ml_prompts p1 p2 = ();
@@ -186,18 +186,18 @@
 (* ML command execution *)
 
 (*Can one redirect 'use' directly to an instream?*)
-fun use_text (tune: string -> string) _ _ _ _ _ txt =
+fun use_text ({tune_source, ...}: use_context) _ _ txt =
   let
     val tmp_name = FileSys.tmpName ();
     val tmp_file = TextIO.openOut tmp_name;
   in
-    TextIO.output (tmp_file, tune txt);
+    TextIO.output (tmp_file, tune_source txt);
     TextIO.closeOut tmp_file;
     use tmp_name;
     FileSys.remove tmp_name
   end;
 
-fun use_file _ _ _ _ _ name = use name;
+fun use_file _ _ name = use name;
 
 
 
--- a/src/Pure/ML-Systems/multithreading.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML-Systems/multithreading.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -21,6 +21,7 @@
   val enabled: unit -> bool
   val no_interrupts: Thread.threadAttribute list
   val regular_interrupts: Thread.threadAttribute list
+  val restricted_interrupts: Thread.threadAttribute list
   val with_attributes: Thread.threadAttribute list ->
     (Thread.threadAttribute list -> 'a -> 'b) -> 'a -> 'b
   val self_critical: unit -> bool
@@ -46,6 +47,9 @@
 val regular_interrupts =
   [Thread.EnableBroadcastInterrupt true, Thread.InterruptState Thread.InterruptAsynchOnce];
 
+val restricted_interrupts =
+  [Thread.EnableBroadcastInterrupt false, Thread.InterruptState Thread.InterruptAsynchOnce];
+
 fun with_attributes _ f x = f [] x;
 
 
--- a/src/Pure/ML-Systems/multithreading_polyml.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML-Systems/multithreading_polyml.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -69,6 +69,9 @@
 val regular_interrupts =
   [Thread.EnableBroadcastInterrupt true, Thread.InterruptState Thread.InterruptAsynchOnce];
 
+val restricted_interrupts =
+  [Thread.EnableBroadcastInterrupt false, Thread.InterruptState Thread.InterruptAsynchOnce];
+
 val safe_interrupts = map
   (fn Thread.InterruptState Thread.InterruptAsynch =>
       Thread.InterruptState Thread.InterruptAsynchOnce
@@ -77,11 +80,9 @@
 fun with_attributes new_atts f x =
   let
     val orig_atts = safe_interrupts (Thread.getAttributes ());
-    fun restore () = Thread.setAttributes orig_atts;
-    val result =
-      (Thread.setAttributes (safe_interrupts new_atts);
-        Exn.Result (f orig_atts x) before restore ())
-      handle e => Exn.Exn e before restore ();
+    val result = Exn.capture (fn () =>
+      (Thread.setAttributes (safe_interrupts new_atts); f orig_atts x)) ();
+    val _ = Thread.setAttributes orig_atts;
   in Exn.release result end;
 
 fun interruptible f = with_attributes regular_interrupts (fn _ => f);
--- a/src/Pure/ML-Systems/polyml-4.1.3.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-4.1.3.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -6,8 +6,10 @@
 use "ML-Systems/polyml_old_basis.ML";
 use "ML-Systems/universal.ML";
 use "ML-Systems/thread_dummy.ML";
+use "ML-Systems/ml_name_space.ML";
 use "ML-Systems/polyml_common.ML";
 use "ML-Systems/polyml_old_compiler4.ML";
+use "ML-Systems/polyml_pp.ML";
 
 val pointer_eq = Address.wordEq;
 
--- a/src/Pure/ML-Systems/polyml-4.1.4.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-4.1.4.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -6,8 +6,10 @@
 use "ML-Systems/polyml_old_basis.ML";
 use "ML-Systems/universal.ML";
 use "ML-Systems/thread_dummy.ML";
+use "ML-Systems/ml_name_space.ML";
 use "ML-Systems/polyml_common.ML";
 use "ML-Systems/polyml_old_compiler4.ML";
+use "ML-Systems/polyml_pp.ML";
 
 val pointer_eq = Address.wordEq;
 
--- a/src/Pure/ML-Systems/polyml-4.2.0.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-4.2.0.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -5,8 +5,10 @@
 
 use "ML-Systems/universal.ML";
 use "ML-Systems/thread_dummy.ML";
+use "ML-Systems/ml_name_space.ML";
 use "ML-Systems/polyml_common.ML";
 use "ML-Systems/polyml_old_compiler4.ML";
+use "ML-Systems/polyml_pp.ML";
 
 val pointer_eq = Address.wordEq;
 
--- a/src/Pure/ML-Systems/polyml-5.0.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-5.0.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -5,8 +5,10 @@
 
 use "ML-Systems/universal.ML";
 use "ML-Systems/thread_dummy.ML";
+use "ML-Systems/ml_name_space.ML";
 use "ML-Systems/polyml_common.ML";
 use "ML-Systems/polyml_old_compiler5.ML";
+use "ML-Systems/polyml_pp.ML";
 
 val pointer_eq = PolyML.pointerEq;
 
--- a/src/Pure/ML-Systems/polyml-5.1.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-5.1.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -4,8 +4,10 @@
 *)
 
 use "ML-Systems/thread_dummy.ML";
+use "ML-Systems/ml_name_space.ML";
 use "ML-Systems/polyml_common.ML";
 use "ML-Systems/polyml_old_compiler5.ML";
+use "ML-Systems/polyml_pp.ML";
 
 val pointer_eq = PolyML.pointerEq;
 
--- a/src/Pure/ML-Systems/polyml-experimental.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-experimental.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -4,6 +4,14 @@
 *)
 
 open Thread;
+
+structure ML_Name_Space =
+struct
+  open PolyML.NameSpace;
+  type T = PolyML.NameSpace.nameSpace;
+  val global = PolyML.globalNameSpace;
+end;
+
 use "ML-Systems/polyml_common.ML";
 use "ML-Systems/multithreading_polyml.ML";
 
@@ -12,21 +20,8 @@
 fun share_common_data () = PolyML.shareCommonData PolyML.rootFunction;
 
 
-(* toplevel pretty printers *)
-
-(*dummy version*)
-fun make_pp path pprint = ();
-fun install_pp _ = ();
-
-
 (* runtime compilation *)
 
-structure ML_NameSpace =
-struct
-  open PolyML.NameSpace;
-  val global = PolyML.globalNameSpace;
-end;
-
 local
 
 fun drop_newline s =
@@ -35,11 +30,11 @@
 
 in
 
-fun use_text (tune: string -> string) str_of_pos
-    name_space (start_line, name) (print, err) verbose txt =
+fun use_text ({tune_source, name_space, str_of_pos, print, error, ...}: use_context)
+    (start_line, name) verbose txt =
   let
     val current_line = ref start_line;
-    val in_buffer = ref (String.explode (tune txt));
+    val in_buffer = ref (String.explode (tune_source txt));
     val out_buffer = ref ([]: string list);
     fun output () = drop_newline (implode (rev (! out_buffer)));
 
@@ -49,10 +44,11 @@
       | c :: cs =>
           (in_buffer := cs; if c = #"\n" then current_line := ! current_line + 1 else (); SOME c));
     fun put s = out_buffer := s :: ! out_buffer;
-    fun put_message {message, hard, location = {startLine = line, ...}, context} =
+    fun put_message {message = msg1, hard, location = {startLine = line, ...}, context} =
      (put (if hard then "Error: " else "Warning: ");
-      PolyML.prettyPrint (put, 76) message;
-      put (str_of_pos line name ^ "\n"));
+      PolyML.prettyPrint (put, 76) msg1;
+      (case context of NONE => () | SOME msg2 => PolyML.prettyPrint (put, 76) msg2);
+      put ("At" ^ str_of_pos line name ^ "\n"));
 
     val parameters =
      [PolyML.Compiler.CPOutStream put,
@@ -65,13 +61,50 @@
         PolyML.compiler (get, parameters) ())
       handle exn =>
        (put ("Exception- " ^ General.exnMessage exn ^ " raised");
-        err (output ()); raise exn);
+        error (output ()); raise exn);
   in if verbose then print (output ()) else () end;
 
-fun use_file tune str_of_pos name_space output verbose name =
+fun use_file context verbose name =
   let
     val instream = TextIO.openIn name;
     val txt = Exn.release (Exn.capture TextIO.inputAll instream before TextIO.closeIn instream);
-  in use_text tune str_of_pos name_space (1, name) output verbose txt end;
+  in use_text context (1, name) verbose txt end;
 
 end;
+
+
+(* toplevel pretty printing *)
+
+val pretty_ml =
+  let
+    fun convert len (PrettyBlock (ind, _, context, prts)) =
+          let
+            fun property name default =
+              (case List.find (fn ContextProperty (a, _) => name = a | _ => false) context of
+                SOME (ContextProperty (_, b)) => b
+              | NONE => default);
+            val bg = property "begin" "";
+            val en = property "end" "";
+            val len' = property "length" len;
+          in ML_Pretty.Block ((bg, en), map (convert len') prts, ind) end
+      | convert len (PrettyString s) =
+          ML_Pretty.String (s, case Int.fromString len of SOME i => i | NONE => size s)
+      | convert _ (PrettyBreak (wd, _)) =
+          ML_Pretty.Break (if wd < 99999 then (false, wd) else (true, 2));
+  in convert "" end;
+
+fun ml_pretty (ML_Pretty.Block ((bg, en), prts, ind)) =
+      let val context =
+        (if bg = "" then [] else [ContextProperty ("begin", bg)]) @
+        (if en = "" then [] else [ContextProperty ("end", en)])
+      in PrettyBlock (ind, false, context, map ml_pretty prts) end
+  | ml_pretty (ML_Pretty.String (s, len)) =
+      if len = size s then PrettyString s
+      else PrettyBlock (0, false, [ContextProperty ("length", Int.toString len)], [PrettyString s])
+  | ml_pretty (ML_Pretty.Break (false, wd)) = PrettyBreak (wd, 0)
+  | ml_pretty (ML_Pretty.Break (true, _)) = PrettyBreak (99999, 0);
+
+fun toplevel_pp context (_: string list) pp =
+  use_text context (1, "pp") false
+    ("addPrettyPrinter (fn _ => fn _ => ml_pretty o Pretty.to_ML o (" ^ pp ^ "))");
+
--- a/src/Pure/ML-Systems/polyml.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML-Systems/polyml.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -4,6 +4,14 @@
 *)
 
 open Thread;
+
+structure ML_Name_Space =
+struct
+  open PolyML.NameSpace;
+  type T = PolyML.NameSpace.nameSpace;
+  val global = PolyML.globalNameSpace;
+end;
+
 use "ML-Systems/polyml_common.ML";
 
 if ml_system = "polyml-5.2"
@@ -17,12 +25,6 @@
 
 (* runtime compilation *)
 
-structure ML_NameSpace =
-struct
-  open PolyML.NameSpace;
-  val global = PolyML.globalNameSpace;
-end;
-
 local
 
 fun drop_newline s =
@@ -31,11 +33,11 @@
 
 in
 
-fun use_text (tune: string -> string) str_of_pos
-    name_space (start_line, name) (print, err) verbose txt =
+fun use_text ({tune_source, name_space, str_of_pos, print, error, ...}: use_context)
+    (start_line, name) verbose txt =
   let
     val current_line = ref start_line;
-    val in_buffer = ref (String.explode (tune txt));
+    val in_buffer = ref (String.explode (tune_source txt));
     val out_buffer = ref ([]: string list);
     fun output () = drop_newline (implode (rev (! out_buffer)));
 
@@ -58,13 +60,16 @@
         PolyML.compiler (get, parameters) ())
       handle exn =>
        (put ("Exception- " ^ General.exnMessage exn ^ " raised");
-        err (output ()); raise exn);
+        error (output ()); raise exn);
   in if verbose then print (output ()) else () end;
 
-fun use_file tune str_of_pos name_space output verbose name =
+fun use_file context verbose name =
   let
     val instream = TextIO.openIn name;
     val txt = Exn.release (Exn.capture TextIO.inputAll instream before TextIO.closeIn instream);
-  in use_text tune str_of_pos name_space (1, name) output verbose txt end;
+  in use_text context (1, name) verbose txt end;
 
 end;
+
+use "ML-Systems/polyml_pp.ML";
+
--- a/src/Pure/ML-Systems/polyml_common.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML-Systems/polyml_common.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -9,7 +9,8 @@
 use "ML-Systems/multithreading.ML";
 use "ML-Systems/time_limit.ML";
 use "ML-Systems/system_shell.ML";
-use "ML-Systems/ml_name_space.ML";
+use "ML-Systems/ml_pretty.ML";
+use "ML-Systems/use_context.ML";
 
 
 (** ML system and platform related **)
@@ -73,13 +74,8 @@
 fun ml_prompts p1 p2 = (PolyML.Compiler.prompt1 := p1; PolyML.Compiler.prompt2 := p2);
 
 
-(* toplevel pretty printing (see also Pure/pure_setup.ML) *)
+(* print depth *)
 
-fun make_pp _ pprint (str, blk, brk, en) _ _ obj =
-  pprint obj (str, fn ind => blk (ind, false), fn wd => brk (wd, 0),
-    fn () => brk (99999, 0), en);
-
-(*print depth*)
 local
   val depth = ref 10;
 in
--- a/src/Pure/ML-Systems/polyml_old_compiler4.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML-Systems/polyml_old_compiler4.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -3,9 +3,9 @@
 Runtime compilation -- for old PolyML.compiler (version 4.x).
 *)
 
-fun use_text (tune: string -> string) _ _ (line: int, name) (print, err) verbose txt =
+fun use_text ({tune_source, print, error, ...}: use_context) (line: int, name) verbose txt =
   let
-    val in_buffer = ref (explode (tune txt));
+    val in_buffer = ref (explode (tune_source txt));
     val out_buffer = ref ([]: string list);
     fun output () = implode (rev (case ! out_buffer of "\n" :: cs => cs | cs => cs));
 
@@ -21,12 +21,12 @@
       | _ => (PolyML.compiler (get, put) (); exec ()));
   in
     exec () handle exn =>
-      (err ((if name = "" then "" else "Error in " ^ name ^ "\n") ^ output ()); raise exn);
+      (error ((if name = "" then "" else "Error in " ^ name ^ "\n") ^ output ()); raise exn);
     if verbose then print (output ()) else ()
   end;
 
-fun use_file tune str_of_pos name_space output verbose name =
+fun use_file context verbose name =
   let
     val instream = TextIO.openIn name;
     val txt = Exn.release (Exn.capture TextIO.inputAll instream before TextIO.closeIn instream);
-  in use_text tune str_of_pos name_space (1, name) output verbose txt end;
+  in use_text context (1, name) verbose txt end;
--- a/src/Pure/ML-Systems/polyml_old_compiler5.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML-Systems/polyml_old_compiler5.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -3,9 +3,9 @@
 Runtime compilation -- for old PolyML.compilerEx (version 5.0, 5.1).
 *)
 
-fun use_text (tune: string -> string) _ _ (line, name) (print, err) verbose txt =
+fun use_text ({tune_source, print, error, ...}: use_context) (line, name) verbose txt =
   let
-    val in_buffer = ref (explode (tune txt));
+    val in_buffer = ref (explode (tune_source txt));
     val out_buffer = ref ([]: string list);
     fun output () = implode (rev (case ! out_buffer of "\n" :: cs => cs | cs => cs));
 
@@ -21,12 +21,12 @@
         [] => ()
       | _ => (PolyML.compilerEx (get, put, fn () => ! current_line, name) (); exec ()));
   in
-    exec () handle exn => (err (output ()); raise exn);
+    exec () handle exn => (error (output ()); raise exn);
     if verbose then print (output ()) else ()
   end;
 
-fun use_file tune str_of_pos name_space output verbose name =
+fun use_file context verbose name =
   let
     val instream = TextIO.openIn name;
     val txt = Exn.release (Exn.capture TextIO.inputAll instream before TextIO.closeIn instream);
-  in use_text tune str_of_pos name_space (1, name) output verbose txt end;
+  in use_text context (1, name) verbose txt end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/polyml_pp.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -0,0 +1,20 @@
+(*  Title:      Pure/ML-Systems/polyml_pp.ML
+
+Toplevel pretty printing for Poly/ML before 5.3.
+*)
+
+fun ml_pprint (print, begin_blk, brk, end_blk) =
+  let
+    fun str "" = ()
+      | str s = print s;
+    fun pprint (ML_Pretty.Block ((bg, en), prts, ind)) =
+          (str bg; begin_blk (ind, false); List.app pprint prts; end_blk (); str en)
+      | pprint (ML_Pretty.String (s, _)) = str s
+      | pprint (ML_Pretty.Break (false, wd)) = brk (wd, 0)
+      | pprint (ML_Pretty.Break (true, _)) = brk (99999, 0);
+  in pprint end;
+
+fun toplevel_pp context (_: string list) pp =
+  use_text context (1, "pp") false
+    ("install_pp (fn args => fn _ => fn _ => ml_pprint args o Pretty.to_ML o (" ^ pp ^ "))");
+
--- a/src/Pure/ML-Systems/smlnj.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML-Systems/smlnj.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -13,6 +13,8 @@
 use "ML-Systems/multithreading.ML";
 use "ML-Systems/system_shell.ML";
 use "ML-Systems/ml_name_space.ML";
+use "ML-Systems/ml_pretty.ML";
+use "ML-Systems/use_context.ML";
 
 
 (*low-level pointer equality*)
@@ -97,25 +99,9 @@
 fun makestring x = "dummy string for SML New Jersey";
 
 
-(* toplevel pretty printing (see also Pure/pure_setup.ML) *)
-
-fun make_pp path pprint =
-  let
-    open PrettyPrint;
-
-    fun pp pps obj =
-      pprint obj
-        (string pps, openHOVBox pps o Rel o dest_int,
-          fn wd => break pps {nsp=dest_int wd, offset=0}, fn () => newline pps,
-          fn () => closeBox pps);
-  in (path, pp) end;
-
-fun install_pp (path, pp) = CompilerPPTable.install_pp path pp;
-
-
 (* ML command execution *)
 
-fun use_text (tune: string -> string) _ _ (line: int, name) (print, err) verbose txt =
+fun use_text ({tune_source, print, error, ...}: use_context) (line, name) verbose txt =
   let
     val ref out_orig = Control.Print.out;
 
@@ -126,22 +112,40 @@
       in String.substring (str, 0, Int.max (0, size str - 1)) end;
   in
     Control.Print.out := out;
-    Backend.Interact.useStream (TextIO.openString (tune txt)) handle exn =>
+    Backend.Interact.useStream (TextIO.openString (tune_source txt)) handle exn =>
       (Control.Print.out := out_orig;
-        err ((if name = "" then "" else "Error in " ^ name ^ "\n") ^ output ()); raise exn);
+        error ((if name = "" then "" else "Error in " ^ name ^ "\n") ^ output ()); raise exn);
     Control.Print.out := out_orig;
     if verbose then print (output ()) else ()
   end;
 
-fun use_file tune str_of_pos name_space output verbose name =
+fun use_file context verbose name =
   let
     val instream = TextIO.openIn name;
     val txt = Exn.release (Exn.capture TextIO.inputAll instream before TextIO.closeIn instream);
-  in use_text tune str_of_pos name_space (1, name) output verbose txt end;
+  in use_text context (1, name) verbose txt end;
+
+fun forget_structure _ = ();
+
+
+(* toplevel pretty printing *)
 
-fun forget_structure name =
-  use_text (fn x => x) (fn _ => "") () (1, "ML") (TextIO.print, fn s => raise Fail s) false
-    ("structure " ^ name ^ " = struct end");
+fun ml_pprint pps =
+  let
+    fun str "" = ()
+      | str s = PrettyPrint.string pps s;
+    fun pprint (ML_Pretty.Block ((bg, en), prts, ind)) =
+          (str bg; PrettyPrint.openHOVBox pps (PrettyPrint.Rel (dest_int ind));
+            List.app pprint prts; PrettyPrint.closeBox pps; str en)
+      | pprint (ML_Pretty.String (s, _)) = str s
+      | pprint (ML_Pretty.Break (false, wd)) = PrettyPrint.break pps {nsp = dest_int wd, offset = 0}
+      | pprint (ML_Pretty.Break (true, _)) = PrettyPrint.newline pps;
+  in pprint end;
+
+fun toplevel_pp context path pp =
+  use_text context (1, "pp") false
+    ("CompilerPPTable.install_pp [" ^ String.concatWith "," (map (fn s => "\"" ^ s ^ "\"")  path) ^
+      "] (fn pps => ml_pprint pps o Pretty.to_ML o (" ^ pp ^ "))");
 
 
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/use_context.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -0,0 +1,13 @@
+(*  Title:      Pure/ML-Systems/use_context.ML
+    Author:     Makarius
+
+Common context for "use" operations (compiler invocation).
+*)
+
+type use_context =
+ {tune_source: string -> string,
+  name_space: ML_Name_Space.T,
+  str_of_pos: int -> string -> string,
+  print: string -> unit,
+  error: string -> unit};
+
--- a/src/Pure/ML/ml_antiquote.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML/ml_antiquote.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -59,7 +59,7 @@
 
 structure P = OuterParse;
 
-val _ = inline "binding"
+val _ = value "binding"
   (Scan.lift (P.position Args.name) >> (fn name => ML_Syntax.atomic (ML_Syntax.make_binding name)));
 
 val _ = value "theory"
--- a/src/Pure/ML/ml_context.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML/ml_context.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -19,8 +19,9 @@
   val the_global_context: unit -> theory
   val the_local_context: unit -> Proof.context
   val exec: (unit -> unit) -> Context.generic -> Context.generic
-  val inherit_env: Proof.context -> Proof.context -> Proof.context
-  val name_space: ML_NameSpace.nameSpace
+  val inherit_env: Context.generic -> Context.generic -> Context.generic
+  val name_space: ML_Name_Space.T
+  val local_context: use_context
   val stored_thms: thm list ref
   val ml_store_thm: string * thm -> unit
   val ml_store_thms: string * thm list -> unit
@@ -29,12 +30,12 @@
       (Proof.context -> string * string) * Proof.context
   val add_antiq: string -> (Position.T -> antiq context_parser) -> unit
   val trace: bool ref
-  val eval_wrapper: (string -> unit) * (string -> 'a) -> bool -> Position.T -> string -> unit
-  val eval: bool -> Position.T -> string -> unit
+  val eval_antiquotes: ML_Lex.token Antiquote.antiquote list * Position.T ->
+    Context.generic option -> (ML_Lex.token list * ML_Lex.token list) * Context.generic option
+  val eval: bool -> Position.T -> Symbol_Pos.text -> unit
   val eval_file: Path.T -> unit
-  val eval_in: Proof.context option -> bool -> Position.T -> string -> unit
-  val evaluate: Proof.context -> (string -> unit) * (string -> 'b) -> bool ->
-    string * (unit -> 'a) option ref -> string -> 'a
+  val eval_in: Proof.context option -> bool -> Position.T -> Symbol_Pos.text -> unit
+  val evaluate: Proof.context -> bool -> string * (unit -> 'a) option ref -> string -> 'a
   val expression: Position.T -> string -> string -> string -> Context.generic -> Context.generic
 end
 
@@ -58,12 +59,12 @@
 structure ML_Env = GenericDataFun
 (
   type T =
-    ML_NameSpace.valueVal Symtab.table *
-    ML_NameSpace.typeVal Symtab.table *
-    ML_NameSpace.fixityVal Symtab.table *
-    ML_NameSpace.structureVal Symtab.table *
-    ML_NameSpace.signatureVal Symtab.table *
-    ML_NameSpace.functorVal Symtab.table;
+    ML_Name_Space.valueVal Symtab.table *
+    ML_Name_Space.typeVal Symtab.table *
+    ML_Name_Space.fixityVal Symtab.table *
+    ML_Name_Space.structureVal Symtab.table *
+    ML_Name_Space.signatureVal Symtab.table *
+    ML_Name_Space.functorVal Symtab.table;
   val empty = (Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty);
   val extend = I;
   fun merge _
@@ -77,25 +78,25 @@
     Symtab.merge (K true) (functor1, functor2));
 );
 
-val inherit_env = Context.proof_map o ML_Env.put o ML_Env.get o Context.Proof;
+val inherit_env = ML_Env.put o ML_Env.get;
 
-val name_space: ML_NameSpace.nameSpace =
+val name_space: ML_Name_Space.T =
   let
     fun lookup sel1 sel2 name =
       Context.thread_data ()
       |> (fn NONE => NONE | SOME context => Symtab.lookup (sel1 (ML_Env.get context)) name)
-      |> (fn NONE => sel2 ML_NameSpace.global name | some => some);
+      |> (fn NONE => sel2 ML_Name_Space.global name | some => some);
 
     fun all sel1 sel2 () =
       Context.thread_data ()
       |> (fn NONE => [] | SOME context => Symtab.dest (sel1 (ML_Env.get context)))
-      |> append (sel2 ML_NameSpace.global ())
+      |> append (sel2 ML_Name_Space.global ())
       |> sort_distinct (string_ord o pairself #1);
 
     fun enter ap1 sel2 entry =
       if is_some (Context.thread_data ()) then
         Context.>> (ML_Env.map (ap1 (Symtab.update entry)))
-      else sel2 ML_NameSpace.global entry;
+      else sel2 ML_Name_Space.global entry;
   in
    {lookupVal    = lookup #1 #lookupVal,
     lookupType   = lookup #2 #lookupType,
@@ -117,6 +118,13 @@
     allFunct     = all #6 #allFunct}
   end;
 
+val local_context: use_context =
+ {tune_source = ML_Parse.fix_ints,
+  name_space = name_space,
+  str_of_pos = Position.str_of oo Position.line_file,
+  print = writeln,
+  error = error};
+
 
 (* theorem bindings *)
 
@@ -131,7 +139,7 @@
       else if not (ML_Syntax.is_identifier name) then
         error ("Cannot bind theorem(s) " ^ quote name ^ " as ML value")
       else setmp stored_thms ths' (fn () =>
-        use_text name_space (0, "") Output.ml_output true
+        use_text local_context (0, "") true
           ("val " ^ name ^ " = " ^ sel ^ "(! ML_Context.stored_thms);")) ();
   in () end;
 
@@ -182,67 +190,70 @@
 local
 
 structure P = OuterParse;
+structure T = OuterLex;
 
 val antiq =
   P.!!! (P.position P.xname -- Args.parse --| Scan.ahead P.eof)
   >> (fn ((x, pos), y) => Args.src ((x, y), pos));
 
 val struct_name = "Isabelle";
+val begin_env = ML_Lex.tokenize ("structure " ^ struct_name ^ " =\nstruct\n");
+val end_env = ML_Lex.tokenize "end;";
 
-fun eval_antiquotes (txt, pos) opt_context =
+in
+
+fun eval_antiquotes (ants, pos) opt_context =
   let
-    val syms = SymbolPos.explode (txt, pos);
-    val ants = Antiquote.read (syms, pos);
     val opt_ctxt = Option.map (Context.Proof o Context.proof_of) opt_context;
     val ((ml_env, ml_body), opt_ctxt') =
-      if not (exists Antiquote.is_antiq ants)
-      then (("", Symbol.escape (SymbolPos.content syms)), opt_ctxt)
+      if forall Antiquote.is_text ants
+      then (([], map (fn Antiquote.Text tok => tok) ants), opt_ctxt)
       else
         let
           val ctxt =
             (case opt_ctxt of
-              NONE => error ("Unknown context -- cannot expand ML antiquotations" ^
-                Position.str_of pos)
+              NONE => error ("No context -- cannot expand ML antiquotations" ^ Position.str_of pos)
             | SOME ctxt => Context.proof_of ctxt);
 
           val lex = #1 (OuterKeyword.get_lexicons ());
-          fun no_decl _ = ("", "");
+          fun no_decl _ = ([], []);
 
-          fun expand (Antiquote.Text s) state = (K ("", Symbol.escape s), state)
-            | expand (Antiquote.Antiq x) (scope, background) =
+          fun expand (Antiquote.Text tok) state = (K ([], [tok]), state)
+            | expand (Antiquote.Antiq (ss, range)) (scope, background) =
                 let
                   val context = Stack.top scope;
-                  val (f, context') = antiquotation (Antiquote.read_antiq lex antiq x) context;
+                  val (f, context') = antiquotation (T.read_antiq lex antiq (ss, #1 range)) context;
                   val (decl, background') = f {background = background, struct_name = struct_name};
-                in (decl, (Stack.map_top (K context') scope, background')) end
+                  val decl' = decl #> pairself (ML_Lex.tokenize #> map (ML_Lex.set_range range));
+                in (decl', (Stack.map_top (K context') scope, background')) end
             | expand (Antiquote.Open _) (scope, background) =
                 (no_decl, (Stack.push scope, background))
             | expand (Antiquote.Close _) (scope, background) =
                 (no_decl, (Stack.pop scope, background));
 
           val (decls, (_, ctxt')) = fold_map expand ants (Stack.init ctxt, ctxt);
-          val ml = decls |> map (fn decl => decl ctxt') |> split_list |> pairself implode;
+          val ml = decls |> map (fn decl => decl ctxt') |> split_list |> pairself flat;
         in (ml, SOME (Context.Proof ctxt')) end;
-  in (("structure " ^ struct_name ^ " =\nstruct\n" ^ ml_env ^ "end;", ml_body), opt_ctxt') end;
-
-in
+  in ((begin_env @ ml_env @ end_env, ml_body), opt_ctxt') end;
 
 val trace = ref false;
 
-fun eval_wrapper pr verbose pos txt =
+fun eval verbose pos txt =
   let
-    fun eval_raw p = use_text name_space
-      (the_default 1 (Position.line_of p), the_default "ML" (Position.file_of p)) pr;
+    fun eval_raw p = use_text local_context
+      (the_default 1 (Position.line_of p), the_default "ML" (Position.file_of p));
 
     (*prepare source text*)
     val _ = Position.report Markup.ML_source pos;
-    val ((env, body), env_ctxt) = eval_antiquotes (txt, pos) (Context.thread_data ());
+    val ants = ML_Lex.read_antiq (Symbol_Pos.explode (txt, pos), pos);
+    val ((env, body), env_ctxt) = eval_antiquotes (ants, pos) (Context.thread_data ())
+      |>> pairself (implode o map ML_Lex.text_of);
     val _ = if ! trace then tracing (cat_lines [env, body]) else ();
 
     (*prepare static ML environment*)
     val _ = Context.setmp_thread_data env_ctxt
         (fn () => (eval_raw Position.none false env; Context.thread_data ())) ()
-      |> (fn NONE => () | SOME context' => Context.>> (ML_Env.put (ML_Env.get context')));
+      |> (fn NONE => () | SOME context' => Context.>> (inherit_env context'));
 
     (*eval ML*)
     val _ = eval_raw pos verbose body;
@@ -254,20 +265,18 @@
 end;
 
 
-(* ML evaluation *)
+(* derived versions *)
 
-val eval = eval_wrapper Output.ml_output;
 fun eval_file path = eval true (Path.position path) (File.read path);
 
 fun eval_in ctxt verbose pos txt =
   Context.setmp_thread_data (Option.map Context.Proof ctxt) (fn () => eval verbose pos txt) ();
 
-fun evaluate ctxt pr verbose (ref_name, r) txt = NAMED_CRITICAL "ML" (fn () =>
+fun evaluate ctxt verbose (ref_name, r) txt = NAMED_CRITICAL "ML" (fn () =>
   let
     val _ = r := NONE;
     val _ = Context.setmp_thread_data (SOME (Context.Proof ctxt)) (fn () =>
-      eval_wrapper pr verbose Position.none
-        ("val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))")) ();
+      eval verbose Position.none ("val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))")) ();
   in (case ! r of NONE => error ("Bad evaluation for " ^ ref_name) | SOME e => e) end) ();
 
 fun expression pos bind body txt =
--- a/src/Pure/ML/ml_lex.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML/ml_lex.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -13,13 +13,18 @@
   val stopper: token Scan.stopper
   val is_regular: token -> bool
   val is_improper: token -> bool
-  val pos_of: token -> string
+  val set_range: Position.range -> token -> token
+  val pos_of: token -> Position.T
   val kind_of: token -> token_kind
   val content_of: token -> string
+  val text_of: token -> string
+  val report_token: token -> unit
   val keywords: string list
   val source: (Symbol.symbol, 'a) Source.source ->
-    (token, (SymbolPos.T, Position.T * (Symbol.symbol, 'a) Source.source)
+    (token, (Symbol_Pos.T, Position.T * (Symbol.symbol, 'a) Source.source)
       Source.source) Source.source
+  val tokenize: string -> token list
+  val read_antiq: Symbol_Pos.T list * Position.T -> token Antiquote.antiquote list
 end;
 
 structure ML_Lex: ML_LEX =
@@ -38,10 +43,10 @@
 
 (* position *)
 
-fun position_of (Token ((pos, _), _)) = pos;
-fun end_position_of (Token ((_, pos), _)) = pos;
+fun set_range range (Token (_, x)) = Token (range, x);
 
-val pos_of = Position.str_of o position_of;
+fun pos_of (Token ((pos, _), _)) = pos;
+fun end_pos_of (Token ((_, pos), _)) = pos;
 
 
 (* control tokens *)
@@ -53,7 +58,7 @@
   | is_eof _ = false;
 
 val stopper =
-  Scan.stopper (fn [] => eof | toks => mk_eof (end_position_of (List.last toks))) is_eof;
+  Scan.stopper (fn [] => eof | toks => mk_eof (end_pos_of (List.last toks))) is_eof;
 
 
 (* token content *)
@@ -63,6 +68,11 @@
 
 fun kind_of (Token (_, (k, _))) = k;
 
+fun text_of tok =
+  (case kind_of tok of
+    Error msg => error msg
+  | _ => Symbol.escape (content_of tok));
+
 fun is_regular (Token (_, (Error _, _))) = false
   | is_regular (Token (_, (EOF, _))) = false
   | is_regular _ = true;
@@ -72,12 +82,33 @@
   | is_improper _ = false;
 
 
+(* markup *)
+
+val token_kind_markup =
+ fn Keyword   => Markup.ML_keyword
+  | Ident     => Markup.ML_ident
+  | LongIdent => Markup.ML_ident
+  | TypeVar   => Markup.ML_tvar
+  | Word      => Markup.ML_numeral
+  | Int       => Markup.ML_numeral
+  | Real      => Markup.ML_numeral
+  | Char      => Markup.ML_char
+  | String    => Markup.ML_string
+  | Space     => Markup.none
+  | Comment   => Markup.ML_comment
+  | Error _   => Markup.ML_malformed
+  | EOF       => Markup.none;
+
+fun report_token (Token ((pos, _), (kind, _))) =
+  Position.report (token_kind_markup kind) pos;
+
+
 
 (** scanners **)
 
-open BasicSymbolPos;
+open Basic_Symbol_Pos;
 
-fun !!! msg = SymbolPos.!!! ("SML lexical error: " ^ msg);
+fun !!! msg = Symbol_Pos.!!! ("SML lexical error: " ^ msg);
 
 
 (* blanks *)
@@ -161,7 +192,8 @@
     Scan.one (Symbol.is_ascii_digit o symbol) >> (fn ((a, b), c) => [a, b, c]);
 
 val scan_str =
-  Scan.one (fn (s, _) => Symbol.is_printable s andalso s <> "\"" andalso s <> "\\") >> single ||
+  Scan.one (fn (s, _) => Symbol.is_regular s andalso s <> "\"" andalso s <> "\\" andalso
+    (not (Symbol.is_char s) orelse Symbol.is_printable s)) >> single ||
   $$$ "\\" @@@ !!! "bad escape character in string" scan_escape;
 
 val scan_gap = $$$ "\\" @@@ scan_blanks1 @@@ $$$ "\\";
@@ -179,17 +211,17 @@
 end;
 
 
-(* token source *)
+(* scan tokens *)
 
 local
 
-fun token k ss = Token (SymbolPos.range ss, (k, SymbolPos.implode ss));
+fun token k ss = Token (Symbol_Pos.range ss, (k, Symbol_Pos.content ss));
 
-val scan = !!! "bad input"
+val scan_ml =
  (scan_char >> token Char ||
   scan_string >> token String ||
   scan_blanks1 >> token Space ||
-  SymbolPos.scan_comment !!! >> token Comment ||
+  Symbol_Pos.scan_comment !!! >> token Comment ||
   Scan.max token_leq
    (scan_keyword >> token Keyword)
    (scan_word >> token Word ||
@@ -199,6 +231,8 @@
     scan_ident >> token Ident ||
     scan_typevar >> token TypeVar));
 
+val scan_antiq = Antiquote.scan || scan_ml >> Antiquote.Text;
+
 fun recover msg =
   Scan.many (((not o Symbol.is_blank) andf Symbol.is_regular) o symbol)
   >> (fn cs => [token (Error msg) cs]);
@@ -206,8 +240,21 @@
 in
 
 fun source src =
-  SymbolPos.source (Position.line 1) src
-  |> Source.source SymbolPos.stopper (Scan.bulk scan) (SOME (false, recover));
+  Symbol_Pos.source (Position.line 1) src
+  |> Source.source Symbol_Pos.stopper (Scan.bulk (!!! "bad input" scan_ml)) (SOME (false, recover));
+
+val tokenize = Source.of_string #> source #> Source.exhaust;
+
+fun read_antiq (syms, pos) =
+  (Source.of_list syms
+    |> Source.source Symbol_Pos.stopper (Scan.bulk (!!! "bad input" scan_antiq))
+      (SOME (false, fn msg => recover msg >> map Antiquote.Text))
+    |> Source.exhaust
+    |> tap (List.app (Antiquote.report report_token))
+    |> tap Antiquote.check_nesting
+    |> tap (List.app (fn Antiquote.Text tok => ignore (text_of tok) | _ => ())))
+  handle ERROR msg =>
+    cat_error msg ("The error(s) above occurred in ML source" ^ Position.str_of pos);
 
 end;
 
--- a/src/Pure/ML/ml_parse.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ML/ml_parse.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -7,38 +7,38 @@
 signature ML_PARSE =
 sig
   val fix_ints: string -> string
+  val global_context: use_context
 end;
 
 structure ML_Parse: ML_PARSE =
 struct
 
-structure T = ML_Lex;
-
-
 (** error handling **)
 
 fun !!! scan =
   let
     fun get_pos [] = " (past end-of-file!)"
-      | get_pos (tok :: _) = T.pos_of tok;
+      | get_pos (tok :: _) = Position.str_of (ML_Lex.pos_of tok);
 
     fun err (toks, NONE) = "SML syntax error" ^ get_pos toks
       | err (toks, SOME msg) = "SML syntax error" ^ get_pos toks ^ ": " ^ msg;
   in Scan.!! err scan end;
 
 fun bad_input x =
-  (Scan.some (fn tok => (case T.kind_of tok of T.Error msg => SOME msg | _ => NONE)) :|--
+  (Scan.some (fn tok => (case ML_Lex.kind_of tok of ML_Lex.Error msg => SOME msg | _ => NONE)) :|--
     (fn msg => Scan.fail_with (K msg))) x;
 
 
 (** basic parsers **)
 
 fun $$$ x =
-  Scan.one (fn tok => T.kind_of tok = T.Keyword andalso T.content_of tok = x) >> T.content_of;
-val int = Scan.one (fn tok => T.kind_of tok = T.Int) >> T.content_of;
+  Scan.one (fn tok => ML_Lex.kind_of tok = ML_Lex.Keyword andalso ML_Lex.content_of tok = x)
+    >> ML_Lex.content_of;
 
-val regular = Scan.one T.is_regular >> T.content_of;
-val improper = Scan.one T.is_improper >> T.content_of;
+val int = Scan.one (fn tok => ML_Lex.kind_of tok = ML_Lex.Int) >> ML_Lex.content_of;
+
+val regular = Scan.one ML_Lex.is_regular >> ML_Lex.content_of;
+val improper = Scan.one ML_Lex.is_improper >> ML_Lex.content_of;
 
 val blanks = Scan.repeat improper >> implode;
 
@@ -55,11 +55,21 @@
 
 fun do_fix_ints s =
   Source.of_string s
-  |> T.source
-  |> Source.source T.stopper (Scan.bulk (!!! fix_int)) NONE
+  |> ML_Lex.source
+  |> Source.source ML_Lex.stopper (Scan.bulk (!!! fix_int)) NONE
   |> Source.exhaust
   |> implode;
 
 val fix_ints = if ml_system_fix_ints then do_fix_ints else I;
 
+
+(* global use_context *)
+
+val global_context: use_context =
+ {tune_source = fix_ints,
+  name_space = ML_Name_Space.global,
+  str_of_pos = Position.str_of oo Position.line_file,
+  print = writeln,
+  error = error};
+
 end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML/ml_test.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -0,0 +1,206 @@
+(*  Title:      Pure/ML/ml_test.ML
+    Author:     Makarius
+
+Test of advanced ML compiler invocation in Poly/ML 5.3.
+*)
+
+signature ML_TEST =
+sig
+  val position_of: Proof.context -> PolyML.location -> Position.T
+  val eval: bool -> Position.T -> ML_Lex.token list -> unit
+end
+
+structure ML_Test: ML_TEST =
+struct
+
+(* extra ML environment *)
+
+structure Extra_Env = GenericDataFun
+(
+  type T = Position.T Inttab.table;  (*position of registered tokens*)
+  val empty = Inttab.empty;
+  val extend = I;
+  fun merge _ = Inttab.merge (K true);
+);
+
+fun inherit_env context =
+  ML_Context.inherit_env context #>
+  Extra_Env.put (Extra_Env.get context);
+
+fun register_tokens toks context =
+  let
+    val regs = map (fn tok => (serial (), tok)) toks;
+    val context' = context
+      |> Extra_Env.map (fold (fn (i, tok) => Inttab.update (i, ML_Lex.pos_of tok)) regs);
+  in (regs, context') end;
+
+fun position_of ctxt
+    ({file, startLine = line, startPosition = i, endPosition = j, ...}: location) =
+  (case pairself (Inttab.lookup (Extra_Env.get (Context.Proof ctxt))) (i, j) of
+    (SOME pos1, SOME pos2) => Position.encode_range (pos1, pos2)
+  | (SOME pos, NONE) => pos
+  | _ => Position.line_file line file);
+
+
+(* parse trees *)
+
+fun report_parse_tree context depth =
+  let
+    val pos_of = position_of (Context.proof_of context);
+    fun report loc (PTtype types) =
+          PolyML.NameSpace.displayTypeExpression (types, depth)
+          |> pretty_ml |> Pretty.from_ML |> Pretty.string_of
+          |> Position.report_text Markup.ML_typing (pos_of loc)
+      | report loc (PTdeclaredAt decl) =
+          Markup.markup (Markup.properties (Position.properties_of (pos_of decl)) Markup.ML_def) ""
+          |> Position.report_text Markup.ML_ref (pos_of loc)
+      | report _ (PTnextSibling tree) = report_tree (tree ())
+      | report _ (PTfirstChild tree) = report_tree (tree ())
+      | report _ _ = ()
+    and report_tree (loc, props) = List.app (report loc) props;
+  in report_tree end;
+
+
+(* eval ML source tokens *)
+
+fun use_text ({name_space = space, print, error, ...}: use_context) verbose pos toks =
+  let
+    (* input *)
+
+    val input = Context.>>> (register_tokens toks);
+    val input_buffer = ref (map (apsnd (String.explode o ML_Lex.text_of)) input);
+
+    val current_line = ref (the_default 1 (Position.line_of pos));
+
+    fun get_index () =
+      the_default 0 (get_first (fn (_, []) => NONE | (i, _) => SOME i) (! input_buffer));
+
+    fun get () =
+      (case ! input_buffer of
+        [] => NONE
+      | (_, []) :: rest => (input_buffer := rest; get ())
+      | (i, c :: cs) :: rest =>
+          (input_buffer := (i, cs) :: rest;
+           if c = #"\n" then current_line := ! current_line + 1 else ();
+           SOME c));
+
+
+    (* output *)
+
+    val output_buffer = ref Buffer.empty;
+    fun output () = Buffer.content (! output_buffer);
+    fun put s = change output_buffer (Buffer.add s);
+
+    fun put_message {message, hard, location, context = _} =
+     (put (if hard then "Error: " else "Warning: ");
+      put (Pretty.string_of (Pretty.from_ML (pretty_ml message)));
+      put (Position.str_of
+        (position_of (Context.proof_of (the (Context.thread_data ()))) location) ^ "\n"));
+
+
+    (* results *)
+
+    val depth = get_print_depth ();
+    val with_struct = ! PolyML.Compiler.printTypesWithStructureName;
+
+    fun apply_result {fixes, types, signatures, structures, functors, values} =
+      let
+        fun add_prefix prefix (PrettyBlock (ind, consistent, context, prts)) =
+              let
+                fun make_prefix context =
+                  (case get_first (fn ContextParentStructure p => SOME p | _ => NONE) context of
+                    SOME (name, sub_context) => make_prefix sub_context ^ name ^ "."
+                  | NONE => prefix);
+                val this_prefix = make_prefix context;
+              in PrettyBlock (ind, consistent, context, map (add_prefix this_prefix) prts) end
+          | add_prefix prefix (prt as PrettyString s) =
+              if prefix = "" then prt else PrettyString (prefix ^ s)
+          | add_prefix _ (prt as PrettyBreak _) = prt;
+
+        fun display disp x =
+          if depth > 0 then
+            (disp x
+              |> with_struct ? add_prefix ""
+              |> pretty_ml |> Pretty.from_ML |> Pretty.string_of |> put; put "\n")
+          else ();
+
+        fun apply_fix (a, b) =
+          (display PolyML.NameSpace.displayFix (a, b); #enterFix space (a, b));
+        fun apply_type (a, b) =
+          (display PolyML.NameSpace.displayType (b, depth); #enterType space (a, b));
+        fun apply_sig (a, b) =
+          (display PolyML.NameSpace.displaySig (b, depth, space); #enterSig space (a, b));
+        fun apply_struct (a, b) =
+          (display PolyML.NameSpace.displayStruct (b, depth, space); #enterStruct space (a, b));
+        fun apply_funct (a, b) =
+          (display PolyML.NameSpace.displayFunct (b, depth, space); #enterFunct space (a, b));
+        fun apply_val (a, b) =
+          (display PolyML.NameSpace.displayVal (b, depth, space); #enterVal space (a, b));
+      in
+        List.app apply_fix fixes;
+        List.app apply_type types;
+        List.app apply_sig signatures;
+        List.app apply_struct structures;
+        List.app apply_funct functors;
+        List.app apply_val values
+      end;
+
+    fun result_fun (phase1, phase2) () =
+     (case phase1 of NONE => ()
+      | SOME parse_tree => report_parse_tree (the (Context.thread_data ())) depth parse_tree;
+      case phase2 of NONE => error "Static Errors"
+      | SOME code => apply_result (Toplevel.program code));
+
+
+    (* compiler invocation *)
+
+    val parameters =
+     [PolyML.Compiler.CPOutStream put,
+      PolyML.Compiler.CPNameSpace space,
+      PolyML.Compiler.CPErrorMessageProc put_message,
+      PolyML.Compiler.CPLineNo (fn () => ! current_line),
+      PolyML.Compiler.CPFileName (the_default "ML" (Position.file_of pos)),
+      PolyML.Compiler.CPLineOffset get_index,
+      PolyML.Compiler.CPCompilerResultFun result_fun];
+    val _ =
+      (while not (List.null (! input_buffer)) do
+        PolyML.compiler (get, parameters) ())
+      handle exn =>
+       (put ("Exception- " ^ General.exnMessage exn ^ " raised");
+        error (output ()); raise exn);
+  in if verbose then print (output ()) else () end;
+
+val eval = use_text ML_Context.local_context;
+
+
+(* ML test command *)
+
+fun ML_test (txt, pos) =
+  let
+    val _ = Position.report Markup.ML_source pos;
+    val ants = ML_Lex.read_antiq (Symbol_Pos.explode (txt, pos), pos);
+    val ((env, body), env_ctxt) = ML_Context.eval_antiquotes (ants, pos) (Context.thread_data ());
+
+    val _ = Context.setmp_thread_data env_ctxt
+        (fn () => (eval false Position.none env; Context.thread_data ())) ()
+      |> (fn NONE => () | SOME context' => Context.>> (inherit_env context'));
+    val _ = eval true pos body;
+    val _ = eval false Position.none (ML_Lex.tokenize "structure Isabelle = struct end");
+  in () end;
+
+
+local structure P = OuterParse and K = OuterKeyword in
+
+fun propagate_env (context as Context.Proof lthy) =
+      Context.Proof (LocalTheory.map_contexts (inherit_env context) lthy)
+  | propagate_env context = context;
+
+val _ =
+  OuterSyntax.command "ML_test" "advanced ML compiler test" (K.tag_ml K.thy_decl)
+    (P.ML_source >> (fn src =>
+      Toplevel.generic_theory (ML_Context.exec (fn () => ML_test src) #> propagate_env)));
+
+end;
+
+end;
+
--- a/src/Pure/Proof/extraction.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Proof/extraction.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -568,7 +568,7 @@
                       (proof_combt
                          (PThm (serial (),
                           ((corr_name name vs', corr_prop, SOME (map TVar (OldTerm.term_tvars corr_prop))),
-                            Future.value (make_proof_body corr_prf))), vfs_of corr_prop))
+                            Future.value (approximate_proof_body corr_prf))), vfs_of corr_prop))
                       (map (get_var_type corr_prop) (vfs_of prop))
                   in
                     ((name, (vs', ((nullt, nullt), (corr_prf, corr_prf')))) :: defs'',
@@ -679,7 +679,7 @@
                       (proof_combt
                         (PThm (serial (),
                          ((corr_name s vs', corr_prop, SOME (map TVar (OldTerm.term_tvars corr_prop))),
-                           Future.value (make_proof_body corr_prf'))), vfs_of corr_prop))
+                           Future.value (approximate_proof_body corr_prf'))), vfs_of corr_prop))
                       (map (get_var_type corr_prop) (vfs_of prop));
                   in
                     ((s, (vs', ((t', u), (corr_prf', corr_prf'')))) :: defs'',
--- a/src/Pure/ProofGeneral/proof_general_emacs.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ProofGeneral/proof_general_emacs.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -76,7 +76,7 @@
 
 fun setup_messages () =
  (Output.writeln_fn := message "" "" "";
-  Output.status_fn := (fn s => ! Output.priority_fn s);
+  Output.status_fn := (fn _ => ());
   Output.priority_fn := message (special "I") (special "J") "";
   Output.tracing_fn := message (special "I" ^ special "V") (special "J") "";
   Output.debug_fn := message (special "K") (special "L") "+++ ";
--- a/src/Pure/ProofGeneral/proof_general_pgip.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ProofGeneral/proof_general_pgip.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -655,13 +655,9 @@
                                   text=[XML.Elem("pgml",[],
                                                  maps YXML.parse_body strings)] })
 
-        fun string_of_thm th = Pretty.string_of
-                                   (Display.pretty_thm_aux
-                                        (Syntax.pp_global (Thm.theory_of_thm th))
-                                        false (* quote *)
-                                        false (* show hyps *)
-                                        [] (* asms *)
-                                        th)
+        fun string_of_thm th = Pretty.string_of (Display.pretty_thm_aux
+            (Syntax.pp_global (Thm.theory_of_thm th))
+            {quote = false, show_hyps = false, show_status = true} [] th)
 
         fun strings_of_thm (thy, name) = map string_of_thm (PureThy.get_thms thy name)
 
--- a/src/Pure/ROOT.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/ROOT.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -101,4 +101,6 @@
 (*configuration for Proof General*)
 cd "ProofGeneral"; use "ROOT.ML"; cd "..";
 
+if ml_system = "polyml-experimental" then use "ML/ml_test.ML" else ();
 use "pure_setup.ML";
+
--- a/src/Pure/Syntax/ast.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Syntax/ast.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -20,7 +20,6 @@
   val str_of_ast: ast -> string
   val pretty_ast: ast -> Pretty.T
   val pretty_rule: ast * ast -> Pretty.T
-  val pprint_ast: ast -> pprint_args -> unit
   val fold_ast: string -> ast list -> ast
   val fold_ast_p: string -> ast list * ast -> ast
   val unfold_ast: string -> ast -> ast list
@@ -79,8 +78,6 @@
   | pretty_ast (Appl asts) =
       Pretty.enclose "(" ")" (Pretty.breaks (map pretty_ast asts));
 
-val pprint_ast = Pretty.pprint o pretty_ast;
-
 fun pretty_rule (lhs, rhs) =
   Pretty.block [pretty_ast lhs, Pretty.str "  ->", Pretty.brk 2, pretty_ast rhs];
 
--- a/src/Pure/Syntax/lexicon.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Syntax/lexicon.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -9,15 +9,15 @@
   val is_identifier: string -> bool
   val is_ascii_identifier: string -> bool
   val is_tid: string -> bool
-  val scan_id: SymbolPos.T list -> SymbolPos.T list * SymbolPos.T list
-  val scan_longid: SymbolPos.T list -> SymbolPos.T list * SymbolPos.T list
-  val scan_tid: SymbolPos.T list -> SymbolPos.T list * SymbolPos.T list
-  val scan_nat: SymbolPos.T list -> SymbolPos.T list * SymbolPos.T list
-  val scan_int: SymbolPos.T list -> SymbolPos.T list * SymbolPos.T list
-  val scan_hex: SymbolPos.T list -> SymbolPos.T list * SymbolPos.T list
-  val scan_bin: SymbolPos.T list -> SymbolPos.T list * SymbolPos.T list
-  val scan_var: SymbolPos.T list -> SymbolPos.T list * SymbolPos.T list
-  val scan_tvar: SymbolPos.T list -> SymbolPos.T list * SymbolPos.T list
+  val scan_id: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
+  val scan_longid: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
+  val scan_tid: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
+  val scan_nat: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
+  val scan_int: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
+  val scan_hex: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
+  val scan_bin: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
+  val scan_var: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
+  val scan_tvar: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
   val implode_xstr: string list -> string
   val explode_xstr: string -> string list
   val read_indexname: string -> indexname
@@ -60,7 +60,7 @@
   val matching_tokens: token * token -> bool
   val valued_token: token -> bool
   val predef_term: string -> token option
-  val tokenize: Scan.lexicon -> bool -> SymbolPos.T list -> token list
+  val tokenize: Scan.lexicon -> bool -> Symbol_Pos.T list -> token list
 end;
 
 structure Lexicon: LEXICON =
@@ -88,9 +88,9 @@
 
 (** basic scanners **)
 
-open BasicSymbolPos;
+open Basic_Symbol_Pos;
 
-fun !!! msg = SymbolPos.!!! ("Inner lexical error: " ^ msg);
+fun !!! msg = Symbol_Pos.!!! ("Inner lexical error: " ^ msg);
 
 val scan_id = Scan.one (Symbol.is_letter o symbol) ::: Scan.many (Symbol.is_letdig o symbol);
 val scan_longid = scan_id @@@ (Scan.repeat1 ($$$ "." @@@ scan_id) >> flat);
@@ -220,7 +220,7 @@
 fun implode_xstr cs = enclose "''" "''" (implode (map (fn "'" => "\\'" | c => c) cs));
 
 fun explode_xstr str =
-  (case Scan.read SymbolPos.stopper scan_str_body (SymbolPos.explode (str, Position.none)) of
+  (case Scan.read Symbol_Pos.stopper scan_str_body (Symbol_Pos.explode (str, Position.none)) of
     SOME cs => map symbol cs
   | _ => error ("Inner lexical error: literal string expected at " ^ quote str));
 
@@ -229,7 +229,7 @@
 (** tokenize **)
 
 fun token_leq (Token (_, s1, _), Token (_, s2, _)) = s1 <= s2;
-fun token kind ss = Token (kind, SymbolPos.content ss, SymbolPos.range ss);
+fun token kind ss = Token (kind, Symbol_Pos.content ss, Symbol_Pos.range ss);
 
 fun tokenize lex xids syms =
   let
@@ -252,16 +252,16 @@
     val scan_lit = Scan.literal lex >> token Literal;
 
     val scan_token =
-      SymbolPos.scan_comment !!! >> token Comment ||
+      Symbol_Pos.scan_comment !!! >> token Comment ||
       Scan.max token_leq scan_lit scan_val ||
       scan_str >> token StrSy ||
       Scan.many1 (Symbol.is_blank o symbol) >> token Space;
   in
     (case Scan.error
-        (Scan.finite SymbolPos.stopper (Scan.repeat scan_token)) syms of
+        (Scan.finite Symbol_Pos.stopper (Scan.repeat scan_token)) syms of
       (toks, []) => toks
-    | (_, ss) => error ("Inner lexical error at: " ^ SymbolPos.content ss ^
-        Position.str_of (#1 (SymbolPos.range ss))))
+    | (_, ss) => error ("Inner lexical error at: " ^ Symbol_Pos.content ss ^
+        Position.str_of (#1 (Symbol_Pos.range ss))))
   end;
 
 
@@ -303,7 +303,7 @@
 (* indexname *)
 
 fun read_indexname s =
-  (case Scan.read SymbolPos.stopper scan_indexname (SymbolPos.explode (s, Position.none)) of
+  (case Scan.read Symbol_Pos.stopper scan_indexname (Symbol_Pos.explode (s, Position.none)) of
     SOME xi => xi
   | _ => error ("Lexical error in variable name: " ^ quote s));
 
@@ -317,16 +317,16 @@
 fun read_var str =
   let
     val scan =
-      $$$ "?" |-- scan_indexname --| Scan.ahead (Scan.one SymbolPos.is_eof) >> var ||
+      $$$ "?" |-- scan_indexname --| Scan.ahead (Scan.one Symbol_Pos.is_eof) >> var ||
       Scan.many (Symbol.is_regular o symbol) >> (free o implode o map symbol);
-  in the (Scan.read SymbolPos.stopper scan (SymbolPos.explode (str, Position.none))) end;
+  in the (Scan.read Symbol_Pos.stopper scan (Symbol_Pos.explode (str, Position.none))) end;
 
 
 (* read_variable *)
 
 fun read_variable str =
   let val scan = $$$ "?" |-- scan_indexname || scan_indexname
-  in Scan.read SymbolPos.stopper scan (SymbolPos.explode (str, Position.none)) end;
+  in Scan.read Symbol_Pos.stopper scan (Symbol_Pos.explode (str, Position.none)) end;
 
 
 (* specific identifiers *)
@@ -341,14 +341,14 @@
 
 fun nat cs =
   Option.map (#1 o Library.read_int o map symbol)
-    (Scan.read SymbolPos.stopper scan_nat cs);
+    (Scan.read Symbol_Pos.stopper scan_nat cs);
 
 in
 
-fun read_nat s = nat (SymbolPos.explode (s, Position.none));
+fun read_nat s = nat (Symbol_Pos.explode (s, Position.none));
 
 fun read_int s =
-  (case SymbolPos.explode (s, Position.none) of
+  (case Symbol_Pos.explode (s, Position.none) of
     ("-", _) :: cs => Option.map ~ (nat cs)
   | cs => nat cs);
 
--- a/src/Pure/Syntax/simple_syntax.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Syntax/simple_syntax.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -21,7 +21,7 @@
 
 fun read scan s =
   (case
-      SymbolPos.explode (s, Position.none) |>
+      Symbol_Pos.explode (s, Position.none) |>
       Lexicon.tokenize lexicon false |>
       filter Lexicon.is_proper |>
       Scan.read Lexicon.stopper scan of
--- a/src/Pure/Syntax/syntax.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Syntax/syntax.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -35,7 +35,7 @@
   val print_trans: syntax -> unit
   val print_syntax: syntax -> unit
   val guess_infix: syntax -> string -> mixfix option
-  val read_token: string -> SymbolPos.T list * Position.T
+  val read_token: string -> Symbol_Pos.T list * Position.T
   val ambiguity_is_error: bool ref
   val ambiguity_level: int ref
   val ambiguity_limit: int ref
@@ -43,12 +43,12 @@
     (((string * int) * sort) list -> string * int -> Term.sort) ->
     (string -> bool * string) -> (string -> string option) ->
     (typ -> typ) -> (sort -> sort) -> Proof.context ->
-    (string -> bool) -> syntax -> typ -> SymbolPos.T list * Position.T -> term
+    (string -> bool) -> syntax -> typ -> Symbol_Pos.T list * Position.T -> term
   val standard_parse_typ: Proof.context -> syntax ->
     ((indexname * sort) list -> indexname -> sort) -> (sort -> sort) ->
-    SymbolPos.T list * Position.T -> typ
+    Symbol_Pos.T list * Position.T -> typ
   val standard_parse_sort: Proof.context -> syntax -> (sort -> sort) ->
-    SymbolPos.T list * Position.T -> sort
+    Symbol_Pos.T list * Position.T -> sort
   datatype 'a trrule =
     ParseRule of 'a * 'a |
     PrintRule of 'a * 'a |
@@ -82,7 +82,7 @@
     (string * string) trrule list -> syntax -> syntax
   val update_trrules_i: ast trrule list -> syntax -> syntax
   val remove_trrules_i: ast trrule list -> syntax -> syntax
-  val parse_token: Markup.T -> string -> SymbolPos.T list * Position.T
+  val parse_token: Markup.T -> string -> Symbol_Pos.T list * Position.T
   val parse_sort: Proof.context -> string -> sort
   val parse_typ: Proof.context -> string -> typ
   val parse_term: Proof.context -> string -> term
@@ -467,7 +467,7 @@
     | XML.Elem ("token", props, []) => ("", Position.of_properties props)
     | XML.Text text => (text, Position.none)
     | _ => (str, Position.none))
-  in (SymbolPos.explode (text, pos), pos) end;
+  in (Symbol_Pos.explode (text, pos), pos) end;
 
 
 (* read_ast *)
--- a/src/Pure/Thy/latex.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Thy/latex.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -88,9 +88,9 @@
 val output_syms = output_symbols o Symbol.explode;
 
 val output_syms_antiq =
-  (fn Antiquote.Text s => output_syms s
+  (fn Antiquote.Text ss => output_symbols (map Symbol_Pos.symbol ss)
     | Antiquote.Antiq (ss, _) =>
-        enclose "%\n\\isaantiq\n" "%\n\\endisaantiq\n" (output_symbols (map SymbolPos.symbol ss))
+        enclose "%\n\\isaantiq\n" "%\n\\endisaantiq\n" (output_symbols (map Symbol_Pos.symbol ss))
     | Antiquote.Open _ => "{\\isaantiqopen}"
     | Antiquote.Close _ => "{\\isaantiqclose}");
 
@@ -117,7 +117,7 @@
     else if T.is_kind T.Verbatim tok then
       let
         val (txt, pos) = T.source_position_of tok;
-        val ants = Antiquote.read (SymbolPos.explode (txt, pos), pos);
+        val ants = Antiquote.read (Symbol_Pos.explode (txt, pos), pos);
         val out = implode (map output_syms_antiq ants);
       in enclose "{\\isacharverbatimopen}" "{\\isacharverbatimclose}" out end
     else output_syms s
--- a/src/Pure/Thy/thy_output.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Thy/thy_output.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -22,7 +22,7 @@
     ({state: Toplevel.state, source: Args.src, context: Proof.context} -> 'a -> string) -> unit
   datatype markup = Markup | MarkupEnv | Verbatim
   val modes: string list ref
-  val eval_antiquote: Scan.lexicon -> Toplevel.state -> SymbolPos.text * Position.T -> string
+  val eval_antiquote: Scan.lexicon -> Toplevel.state -> Symbol_Pos.text * Position.T -> string
   val present_thy: Scan.lexicon -> (string -> string list) -> (markup -> string -> bool) ->
     (Toplevel.transition * Toplevel.state) list -> (OuterLex.token, 'a) Source.source -> Buffer.T
   val pretty_text: string -> Pretty.T
@@ -147,18 +147,18 @@
 
 fun eval_antiquote lex state (txt, pos) =
   let
-    fun expand (Antiquote.Text s) = s
-      | expand (Antiquote.Antiq x) =
-          let val (opts, src) = Antiquote.read_antiq lex antiq x in
+    fun expand (Antiquote.Text ss) = Symbol_Pos.content ss
+      | expand (Antiquote.Antiq (ss, (pos, _))) =
+          let val (opts, src) = T.read_antiq lex antiq (ss, pos) in
             options opts (fn () => command src state) ();  (*preview errors!*)
             PrintMode.with_modes (! modes @ Latex.modes)
               (Output.no_warnings (options opts (fn () => command src state))) ()
           end
       | expand (Antiquote.Open _) = ""
       | expand (Antiquote.Close _) = "";
-    val ants = Antiquote.read (SymbolPos.explode (txt, pos), pos);
+    val ants = Antiquote.read (Symbol_Pos.explode (txt, pos), pos);
   in
-    if Toplevel.is_toplevel state andalso exists Antiquote.is_antiq ants then
+    if Toplevel.is_toplevel state andalso not (forall Antiquote.is_text ants) then
       error ("Unknown context -- cannot expand document antiquotations" ^ Position.str_of pos)
     else implode (map expand ants)
   end;
@@ -577,7 +577,7 @@
 fun ml_text name ml = antiquotation name (Scan.lift Args.name_source_position)
   (fn {context, ...} => fn (txt, pos) =>
    (ML_Context.eval_in (SOME context) false pos (ml txt);
-    SymbolPos.content (SymbolPos.explode (txt, pos))
+    Symbol_Pos.content (Symbol_Pos.explode (txt, pos))
     |> (if ! quotes then quote else I)
     |> (if ! display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"
     else
--- a/src/Pure/Thy/thy_syntax.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Thy/thy_syntax.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -7,7 +7,7 @@
 signature THY_SYNTAX =
 sig
   val token_source: Scan.lexicon * Scan.lexicon -> Position.T -> (string, 'a) Source.source ->
-    (OuterLex.token, (SymbolPos.T, Position.T * (Symbol.symbol, (string, 'a)
+    (OuterLex.token, (Symbol_Pos.T, Position.T * (Symbol.symbol, (string, 'a)
       Source.source) Source.source) Source.source) Source.source
   val parse_tokens: Scan.lexicon * Scan.lexicon -> Position.T -> string -> OuterLex.token list
   val present_token: OuterLex.token -> output
--- a/src/Pure/Tools/find_theorems.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Tools/find_theorems.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -336,7 +336,9 @@
 
 fun find_theorems ctxt opt_goal rem_dups raw_criteria =
   let
-    val add_prems = Seq.hd o (TRY (Method.insert_tac (Assumption.all_prems_of ctxt) 1));
+    val assms = ProofContext.get_fact ctxt (Facts.named "local.assms")
+                handle ERROR _ => [];
+    val add_prems = Seq.hd o (TRY (Method.insert_tac assms 1));
     val opt_goal' = Option.map add_prems opt_goal;
 
     val criteria = map (apsnd (read_criterion ctxt)) raw_criteria;
--- a/src/Pure/Tools/xml_syntax.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/Tools/xml_syntax.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -158,7 +158,7 @@
   | proof_of_xml (XML.Elem ("PThm", [("name", s)], term :: opttypes)) =
       (* FIXME? *)
       PThm (serial (), ((s, term_of_xml term, opttypes_of_xml opttypes),
-        Future.value (Proofterm.make_proof_body MinProof)))
+        Future.value (Proofterm.approximate_proof_body MinProof)))
   | proof_of_xml (XML.Elem ("PAxm", [("name", s)], term :: opttypes)) =
       PAxm (s, term_of_xml term, opttypes_of_xml opttypes)
   | proof_of_xml (XML.Elem ("Oracle", [("name", s)], term :: opttypes)) =
--- a/src/Pure/context.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/context.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -27,8 +27,6 @@
   val display_names: theory -> string list
   val pretty_thy: theory -> Pretty.T
   val string_of_thy: theory -> string
-  val pprint_thy: theory -> pprint_args -> unit
-  val pprint_thy_ref: theory_ref -> pprint_args -> unit
   val pretty_abbrev_thy: theory -> Pretty.T
   val str_of_thy: theory -> string
   val deref: theory_ref -> theory
@@ -228,7 +226,6 @@
 
 val pretty_thy = Pretty.str_list "{" "}" o display_names;
 val string_of_thy = Pretty.string_of o pretty_thy;
-val pprint_thy = Pretty.pprint o pretty_thy;
 
 fun pretty_abbrev_thy thy =
   let
@@ -256,8 +253,6 @@
     else thy_ref
   end;
 
-val pprint_thy_ref = Pretty.pprint o pretty_thy o deref;
-
 
 (* build ids *)
 
--- a/src/Pure/display.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/display.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -17,7 +17,8 @@
 sig
   include BASIC_DISPLAY
   val pretty_flexpair: Pretty.pp -> term * term -> Pretty.T
-  val pretty_thm_aux: Pretty.pp -> bool -> bool -> term list -> thm -> Pretty.T
+  val pretty_thm_aux: Pretty.pp -> {quote: bool, show_hyps: bool, show_status: bool} ->
+    term list -> thm -> Pretty.T
   val pretty_thm: thm -> Pretty.T
   val string_of_thm: thm -> string
   val pretty_thms: thm list -> Pretty.T
@@ -57,7 +58,20 @@
 fun pretty_flexpair pp (t, u) = Pretty.block
   [Pretty.term pp t, Pretty.str " =?=", Pretty.brk 1, Pretty.term pp u];
 
-fun pretty_thm_aux pp quote show_hyps' asms raw_th =
+fun display_status false _ = ""
+  | display_status true th =
+      let
+        val {oracle = oracle0, unfinished, failed} = Thm.status_of th;
+        val oracle = oracle0 andalso (not (! quick_and_dirty) orelse ! show_hyps);
+      in
+        if failed then "!!"
+        else if oracle andalso unfinished then "!?"
+        else if oracle then "!"
+        else if unfinished then "?"
+        else ""
+      end;
+
+fun pretty_thm_aux pp {quote, show_hyps = show_hyps', show_status} asms raw_th =
   let
     val th = Thm.strip_shyps raw_th;
     val {hyps, tpairs, prop, ...} = Thm.rep_thm th;
@@ -68,27 +82,25 @@
     val prt_term = q o Pretty.term pp;
 
     val hyps' = if ! show_hyps then hyps else subtract (op aconv) asms hyps;
-(* FIXME
-    val ora' = Thm.oracle_of th andalso (! show_hyps orelse not (! quick_and_dirty)); *)
-    val ora' = false;
+    val status = display_status show_status th;
 
     val hlen = length xshyps + length hyps' + length tpairs;
     val hsymbs =
-      if hlen = 0 andalso not ora' then []
+      if hlen = 0 andalso status = "" then []
       else if ! show_hyps orelse show_hyps' then
         [Pretty.brk 2, Pretty.list "[" "]"
           (map (q o pretty_flexpair pp) tpairs @ map prt_term hyps' @
            map (Pretty.sort pp) xshyps @
-            (if ora' then [Pretty.str "!"] else []))]
-      else [Pretty.brk 2, Pretty.str ("[" ^ implode (replicate hlen ".") ^
-        (if ora' then "!" else "") ^ "]")];
+            (if status = "" then [] else [Pretty.str status]))]
+      else [Pretty.brk 2, Pretty.str ("[" ^ implode (replicate hlen ".") ^ status ^ "]")];
     val tsymbs =
       if null tags orelse not (! show_tags) then []
       else [Pretty.brk 1, pretty_tags tags];
   in Pretty.block (prt_term prop :: (hsymbs @ tsymbs)) end;
 
 fun pretty_thm th =
-  pretty_thm_aux (Syntax.pp_global (Thm.theory_of_thm th)) true false [] th;
+  pretty_thm_aux (Syntax.pp_global (Thm.theory_of_thm th))
+    {quote = true, show_hyps = false, show_status = true} [] th;
 
 val string_of_thm = Pretty.string_of o pretty_thm;
 
--- a/src/Pure/mk	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/mk	Sat Mar 28 00:13:01 2009 +0100
@@ -91,6 +91,8 @@
     -e "val ml_system = \"$ML_SYSTEM\";" \
     -e "val ml_platform = \"$ML_PLATFORM\";" \
     -e "use\"$COMPAT\" handle _ => exit 1;" \
+    -e "structure Isar = struct fun main () = () end;" \
+    -e "ml_prompts \"ML> \" \"ML# \";" \
     -q -w RAW_ML_SYSTEM RAW > "$LOG" 2>&1
   RC="$?"
 elif [ -n "$RAW_SESSION" ]; then
@@ -111,6 +113,7 @@
     -e "val ml_system = \"$ML_SYSTEM\";" \
     -e "val ml_platform = \"$ML_PLATFORM\";" \
     -e "(use\"$COMPAT\"; use\"ROOT.ML\") handle _ => exit 1;" \
+    -e "ml_prompts \"ML> \" \"ML# \";" \
     -f -c -q -w RAW_ML_SYSTEM Pure > "$LOG" 2>&1
   RC="$?"
 fi
--- a/src/Pure/name.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/name.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -8,6 +8,7 @@
 sig
   val uu: string
   val aT: string
+  val of_binding: binding -> string
   val bound: int -> string
   val is_bound: string -> bool
   val internal: string -> string
@@ -41,6 +42,11 @@
 
 (** special variable names **)
 
+(* checked binding *)
+
+val of_binding = Long_Name.base_name o NameSpace.full_name NameSpace.default_naming;
+
+
 (* encoded bounds *)
 
 (*names for numbered variables --
--- a/src/Pure/proofterm.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/proofterm.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -40,16 +40,16 @@
     {oracles: oracle OrdList.T, thms: pthm OrdList.T, proof: proof}
   val join_proof: proof_body future -> proof
   val proof_of: proof_body -> proof
+  val fold_proof_atoms: bool -> (proof -> 'a -> 'a) -> proof list -> 'a -> 'a
   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
+  val status_of: proof_body list -> {failed: bool, oracle: bool, unfinished: bool}
 
   val oracle_ord: oracle * oracle -> order
   val thm_ord: pthm * pthm -> order
-  val make_proof_body: proof -> proof_body
   val merge_oracles: oracle OrdList.T -> oracle OrdList.T -> oracle OrdList.T
-  val make_oracles: proof -> oracle OrdList.T
   val merge_thms: pthm OrdList.T -> pthm OrdList.T -> pthm OrdList.T
-  val make_thms: proof -> pthm OrdList.T
+  val all_oracles_of: proof_body -> oracle OrdList.T
+  val approximate_proof_body: proof -> proof_body
 
   (** primitive operations **)
   val proof_combt: proof * term list -> proof
@@ -106,11 +106,11 @@
   val equal_intr: term -> term -> proof -> proof -> proof
   val equal_elim: term -> term -> proof -> proof -> proof
   val axm_proof: string -> term -> proof
-  val oracle_proof: string -> term -> proof
+  val oracle_proof: string -> term -> oracle * proof
   val promise_proof: theory -> serial -> term -> proof
-  val fulfill_proof: theory -> (serial * proof) list -> proof_body -> proof_body
+  val fulfill_proof: theory -> (serial * proof_body) list -> proof_body -> proof_body
   val thm_proof: theory -> string -> term list -> term ->
-    (serial * proof future) list -> proof_body -> pthm * proof
+    (serial * proof_body future) list -> proof_body -> pthm * proof
   val get_name: term list -> term -> proof -> string
 
   (** rewriting on proof terms **)
@@ -159,17 +159,6 @@
 
 (***** proof atoms *****)
 
-fun fold_body_thms f =
-  let
-    fun app (PBody {thms, ...}) = thms |> fold (fn (i, (name, prop, body)) => fn (x, seen) =>
-      if Inttab.defined seen i then (x, seen)
-      else
-        let
-          val body' = Future.join body;
-          val (x', seen') = app body' (x, Inttab.update (i, ()) seen);
-        in (f (name, prop, body') x', seen') end);
-  in fn bodies => fn x => #1 (fold app bodies (x, Inttab.empty)) end;
-
 fun fold_proof_atoms all f =
   let
     fun app (Abst (_, _, prf)) = app prf
@@ -185,32 +174,71 @@
       | app prf = (fn (x, seen) => (f prf x, seen));
   in fn prfs => fn x => #1 (fold app prfs (x, Inttab.empty)) end;
 
+fun fold_body_thms f =
+  let
+    fun app (PBody {thms, ...}) = thms |> fold (fn (i, (name, prop, body)) => fn (x, seen) =>
+      if Inttab.defined seen i then (x, seen)
+      else
+        let
+          val body' = Future.join body;
+          val (x', seen') = app body' (x, Inttab.update (i, ()) seen);
+        in (f (name, prop, body') x', seen') end);
+  in fn bodies => fn x => #1 (fold app bodies (x, Inttab.empty)) end;
+
+fun status_of bodies =
+  let
+    fun status (PBody {oracles, thms, ...}) x =
+      let
+        val ((oracle, unfinished, failed), seen) =
+          (thms, x) |-> fold (fn (i, (_, _, body)) => fn (st, seen) =>
+            if Inttab.defined seen i then (st, seen)
+            else
+              let val seen' = Inttab.update (i, ()) seen in
+                (case Future.peek body of
+                  SOME (Exn.Result body') => status body' (st, seen')
+                | SOME (Exn.Exn _) =>
+                    let val (oracle, unfinished, _) = st
+                    in ((oracle, unfinished, true), seen') end
+                | NONE =>
+                    let val (oracle, _, failed) = st
+                    in ((oracle, true, failed), seen') end)
+              end);
+      in ((oracle orelse not (null oracles), unfinished, failed), seen) end;
+    val (oracle, unfinished, failed) = #1 (fold status bodies ((false, false, false), Inttab.empty));
+  in {oracle = oracle, unfinished = unfinished, failed = failed} end;
+
 
 (* proof body *)
 
 val oracle_ord = prod_ord fast_string_ord TermOrd.fast_term_ord;
 fun thm_ord ((i, _): pthm, (j, _)) = int_ord (j, i);
 
-fun make_body prf =
+val merge_oracles = OrdList.union oracle_ord;
+val merge_thms = OrdList.union thm_ord;
+
+val all_oracles_of =
+  let
+    fun collect (PBody {oracles, thms, ...}) = thms |> fold (fn (i, (_, _, body)) => fn (x, seen) =>
+      if Inttab.defined seen i then (x, seen)
+      else
+        let
+          val body' = Future.join body;
+          val (x', seen') = collect body' (x, Inttab.update (i, ()) seen);
+        in (merge_oracles oracles x', seen') end);
+  in fn body => #1 (collect body ([], Inttab.empty)) end;
+
+fun approximate_proof_body prf =
   let
     val (oracles, thms) = fold_proof_atoms false
       (fn Oracle (s, prop, _) => apfst (cons (s, prop))
         | PThm (i, ((name, prop, _), body)) => apsnd (cons (i, (name, prop, body)))
         | _ => I) [prf] ([], []);
-  in (OrdList.make oracle_ord oracles, OrdList.make thm_ord thms) end;
-
-fun make_proof_body prf =
-  let val (oracles, thms) = make_body prf
-  in PBody {oracles = oracles, thms = thms, proof = prf} end;
-
-val make_oracles = #1 o make_body;
-val make_thms = #2 o make_body;
-
-val merge_oracles = OrdList.union oracle_ord;
-val merge_thms = OrdList.union thm_ord;
-
-fun merge_body (oracles1, thms1) (oracles2, thms2) =
-  (merge_oracles oracles1 oracles2, merge_thms thms1 thms2);
+  in
+    PBody
+     {oracles = OrdList.make oracle_ord oracles,
+      thms = OrdList.make thm_ord thms,
+      proof = prf}
+  end;
 
 
 (***** proof objects with different levels of detail *****)
@@ -907,8 +935,8 @@
 val dummy = Const (Term.dummy_patternN, dummyT);
 
 fun oracle_proof name prop =
-  if !proofs = 0 then Oracle (name, dummy, NONE)
-  else gen_axm_proof Oracle name prop;
+  if ! proofs = 0 then ((name, dummy), Oracle (name, dummy, NONE))
+  else ((name, prop), gen_axm_proof Oracle name prop);
 
 fun shrink_proof thy =
   let
@@ -1212,16 +1240,17 @@
   in Promise (i, prop, map TVar (Term.add_tvars prop [])) end;
 
 fun fulfill_proof _ [] body0 = body0
-  | fulfill_proof thy promises body0 =
+  | fulfill_proof thy ps body0 =
       let
         val PBody {oracles = oracles0, thms = thms0, proof = proof0} = body0;
-        val (oracles, thms) = fold (merge_body o make_body o #2) promises (oracles0, thms0);
+        val oracles = fold (fn (_, PBody {oracles, ...}) => merge_oracles oracles) ps oracles0;
+        val thms = fold (fn (_, PBody {thms, ...}) => merge_thms thms) ps thms0;
+        val proofs = fold (fn (i, PBody {proof, ...}) => Inttab.update (i, proof)) ps Inttab.empty;
 
-        val tab = Inttab.make promises;
         fun fill (Promise (i, prop, Ts)) =
-            (case Inttab.lookup tab i of
+            (case Inttab.lookup proofs i of
               NONE => NONE
-            | SOME p => SOME (instantiate (Term.add_tvars prop [] ~~ Ts, []) p))
+            | SOME prf => SOME (instantiate (Term.add_tvars prop [] ~~ Ts, []) prf))
           | fill _ = NONE;
         val (rules, procs) = get_data thy;
         val proof = rewrite_prf fst (rules, K fill :: procs) proof0;
--- a/src/Pure/pure_setup.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/pure_setup.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -27,29 +27,30 @@
 
 (* ML toplevel pretty printing *)
 
-install_pp (make_pp ["Task_Queue", "task"] (Pretty.pprint o Pretty.str o Task_Queue.str_of_task));
-install_pp (make_pp ["Task_Queue", "group"] (Pretty.pprint o Pretty.str o Task_Queue.str_of_group));
-install_pp (make_pp ["Position", "T"] (Pretty.pprint o Pretty.enum "," "{" "}" o
-  map (fn (x, y) => Pretty.str (x ^ "=" ^ y)) o Position.properties_of));
-install_pp (make_pp ["Thm", "thm"] ProofDisplay.pprint_thm);
-install_pp (make_pp ["Thm", "cterm"] ProofDisplay.pprint_cterm);
-install_pp (make_pp ["Binding", "binding"] (Pretty.pprint o Pretty.str o quote o Binding.str_of));
-install_pp (make_pp ["Thm", "ctyp"] ProofDisplay.pprint_ctyp);
-install_pp (make_pp ["Context", "theory"] Context.pprint_thy);
-install_pp (make_pp ["Context", "theory_ref"] Context.pprint_thy_ref);
-install_pp (make_pp ["Context", "proof"] ProofDisplay.pprint_context);
-install_pp (make_pp ["Syntax", "ast"] Syntax.pprint_ast);
-install_pp (make_pp ["typ"] (ProofDisplay.pprint_typ Pure.thy));
-install_pp (make_pp ["Path", "T"] (Pretty.pprint o Pretty.str o quote o Path.implode));
-install_pp (make_pp ["File", "ident"] (Pretty.pprint o Pretty.str o quote o File.rep_ident));
+toplevel_pp ["Task_Queue", "task"] "Pretty.str o Task_Queue.str_of_task";
+toplevel_pp ["Task_Queue", "group"] "Pretty.str o Task_Queue.str_of_group";
+toplevel_pp ["Position", "T"] "Pretty.position";
+toplevel_pp ["Binding", "binding"] "Pretty.str o quote o Binding.str_of";
+toplevel_pp ["Thm", "thm"] "ProofDisplay.pp_thm";
+toplevel_pp ["Thm", "cterm"] "ProofDisplay.pp_cterm";
+toplevel_pp ["Thm", "ctyp"] "ProofDisplay.pp_ctyp";
+toplevel_pp ["typ"] "ProofDisplay.pp_typ Pure.thy";
+toplevel_pp ["Context", "theory"] "Context.pretty_thy";
+toplevel_pp ["Context", "theory_ref"] "Context.pretty_thy o Theory.deref";
+toplevel_pp ["Context", "proof"] "ProofDisplay.pp_context";
+toplevel_pp ["Syntax", "ast"] "Syntax.pretty_ast";
+toplevel_pp ["Path", "T"] "Pretty.str o quote o Path.implode";
+toplevel_pp ["File", "ident"] "Pretty.str o quote o File.rep_ident";
 
-if String.isPrefix "polyml" ml_system then use "ML-Systems/install_pp_polyml.ML"
+if File.exists (Path.explode ("ML-Systems/install_pp_" ^ ml_system ^ ".ML"))
+then use ("ML-Systems/install_pp_" ^ ml_system ^ ".ML")
+else if String.isPrefix "polyml" ml_system
+then use "ML-Systems/install_pp_polyml.ML"
 else ();
 
 
 (* misc *)
 
 val cd = File.cd o Path.explode;
-ml_prompts "ML> " "ML# ";
 
 Proofterm.proofs := 0;
--- a/src/Pure/sign.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/sign.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -434,7 +434,7 @@
 
 fun add_types types thy = thy |> map_sign (fn (naming, syn, tsig, consts) =>
   let
-    val syn' = Syntax.update_type_gram (map (fn (a, n, mx) => (Binding.name_of a, n, mx)) types) syn;
+    val syn' = Syntax.update_type_gram (map (fn (a, n, mx) => (Name.of_binding a, n, mx)) types) syn;
     val decls = map (fn (a, n, mx) => (Binding.map_name (Syntax.type_name mx) a, n)) types;
     val tags = [(Markup.theory_nameN, Context.theory_name thy)];
     val tsig' = fold (Type.add_type naming tags) decls tsig;
@@ -445,7 +445,7 @@
 
 fun add_nonterminals ns thy = thy |> map_sign (fn (naming, syn, tsig, consts) =>
   let
-    val syn' = Syntax.update_consts (map Binding.name_of ns) syn;
+    val syn' = Syntax.update_consts (map Name.of_binding ns) syn;
     val tsig' = fold (Type.add_nonterminal naming []) ns tsig;
   in (naming, syn', tsig', consts) end);
 
@@ -456,7 +456,7 @@
   thy |> map_sign (fn (naming, syn, tsig, consts) =>
     let
       val ctxt = ProofContext.init thy;
-      val syn' = Syntax.update_type_gram [(Binding.name_of a, length vs, mx)] syn;
+      val syn' = Syntax.update_type_gram [(Name.of_binding a, length vs, mx)] syn;
       val b = Binding.map_name (Syntax.type_name mx) a;
       val abbr = (b, vs, certify_typ_mode Type.mode_syntax thy (parse_typ ctxt rhs))
         handle ERROR msg => cat_error msg ("in type abbreviation " ^ quote (Binding.str_of b));
@@ -504,10 +504,10 @@
     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 (Binding.name_of raw_b) raw_mx;
+        val (mx_name, mx) = Syntax.const_mixfix (Name.of_binding raw_b) raw_mx;
         val b = Binding.map_name (K mx_name) raw_b;
         val c = full_name thy b;
-        val c_syn = if authentic then Syntax.constN ^ c else Binding.name_of b;
+        val c_syn = if authentic then Syntax.constN ^ c else Name.of_binding b;
         val T = (prepT raw_T handle TYPE (msg, _, _) => error msg) handle ERROR msg =>
           cat_error msg ("in declaration of constant " ^ quote (Binding.str_of b));
         val T' = Logic.varifyT T;
@@ -568,7 +568,7 @@
 fun primitive_class (bclass, classes) thy =
   thy |> map_sign (fn (naming, syn, tsig, consts) =>
     let
-      val syn' = Syntax.update_consts [Binding.name_of bclass] syn;
+      val syn' = Syntax.update_consts [Name.of_binding bclass] syn;
       val tsig' = Type.add_class (Syntax.pp_global thy) naming (bclass, classes) tsig;
     in (naming, syn', tsig', consts) end)
   |> add_consts_i [(Binding.map_name Logic.const_of_class bclass, Term.a_itselfT --> propT, NoSyn)];
--- a/src/Pure/simplifier.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/simplifier.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -10,15 +10,8 @@
   include BASIC_META_SIMPLIFIER
   val change_simpset: (simpset -> simpset) -> unit
   val simpset_of: theory -> simpset
-  val simpset: unit -> simpset
-  val SIMPSET: (simpset -> tactic) -> tactic
-  val SIMPSET': (simpset -> 'a -> tactic) -> 'a -> tactic
-  val Addsimps: thm list -> unit
-  val Delsimps: thm list -> unit
   val Addsimprocs: simproc list -> unit
   val Delsimprocs: simproc list -> unit
-  val Addcongs: thm list -> unit
-  val Delcongs: thm list -> unit
   val local_simpset_of: Proof.context -> simpset
   val generic_simp_tac: bool -> bool * bool * bool -> simpset -> int -> tactic
   val safe_asm_full_simp_tac: simpset -> int -> tactic
@@ -27,11 +20,6 @@
   val          full_simp_tac: simpset -> int -> tactic
   val        asm_lr_simp_tac: simpset -> int -> tactic
   val      asm_full_simp_tac: simpset -> int -> tactic
-  val               Simp_tac:            int -> tactic
-  val           Asm_simp_tac:            int -> tactic
-  val          Full_simp_tac:            int -> tactic
-  val        Asm_lr_simp_tac:            int -> tactic
-  val      Asm_full_simp_tac:            int -> tactic
   val          simplify: simpset -> thm -> thm
   val      asm_simplify: simpset -> thm -> thm
   val     full_simplify: simpset -> thm -> thm
@@ -138,17 +126,9 @@
 fun map_simpset f = Context.theory_map (map_ss f);
 fun change_simpset f = Context.>> (Context.map_theory (map_simpset f));
 fun simpset_of thy = MetaSimplifier.context (ProofContext.init thy) (get_ss (Context.Theory thy));
-val simpset = simpset_of o ML_Context.the_global_context;
 
-fun SIMPSET tacf st = tacf (simpset_of (Thm.theory_of_thm st)) st;
-fun SIMPSET' tacf i st = tacf (simpset_of (Thm.theory_of_thm st)) i st;
-
-fun Addsimps args = change_simpset (fn ss => ss addsimps args);
-fun Delsimps args = change_simpset (fn ss => ss delsimps args);
 fun Addsimprocs args = change_simpset (fn ss => ss addsimprocs args);
 fun Delsimprocs args = change_simpset (fn ss => ss delsimprocs args);
-fun Addcongs args = change_simpset (fn ss => ss addcongs args);
-fun Delcongs args = change_simpset (fn ss => ss delcongs args);
 
 
 (* local simpset *)
@@ -283,13 +263,6 @@
 val asm_full_simp_tac = generic_simp_tac false (true, true, true);
 val safe_asm_full_simp_tac = generic_simp_tac true (true, true, true);
 
-(*the abstraction over the proof state delays the dereferencing*)
-fun          Simp_tac i st =          simp_tac (simpset ()) i st;
-fun      Asm_simp_tac i st =      asm_simp_tac (simpset ()) i st;
-fun     Full_simp_tac i st =     full_simp_tac (simpset ()) i st;
-fun   Asm_lr_simp_tac i st =   asm_lr_simp_tac (simpset ()) i st;
-fun Asm_full_simp_tac i st = asm_full_simp_tac (simpset ()) i st;
-
 
 (* conversions *)
 
--- a/src/Pure/thm.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Pure/thm.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -144,6 +144,7 @@
   val freezeT: thm -> thm
   val future: thm future -> cterm -> thm
   val pending_groups: thm -> Task_Queue.group list -> Task_Queue.group list
+  val status_of: thm -> {oracle: bool, unfinished: bool, failed: bool}
   val proof_body_of: thm -> proof_body
   val proof_of: thm -> proof
   val join_proof: thm -> unit
@@ -1635,20 +1636,33 @@
   end;
 
 
-(* pending task groups *)
+(* derivation status *)
+
+fun raw_proof_body_of (Thm (Deriv {body, ...}, _)) = body;
+val raw_proof_of = Proofterm.proof_of o raw_proof_body_of;
 
 fun pending_groups (Thm (Deriv {open_promises, ...}, _)) =
   fold (insert Task_Queue.eq_group o Future.group_of o #2) open_promises;
 
+fun status_of (Thm (Deriv {promises, body, ...}, _)) =
+  let
+    val ps = map (Future.peek o snd) promises;
+    val bodies = body ::
+      map_filter (fn SOME (Exn.Result th) => SOME (raw_proof_body_of th) | _ => NONE) ps;
+    val {oracle, unfinished, failed} = Pt.status_of bodies;
+  in
+   {oracle = oracle,
+    unfinished = unfinished orelse exists is_none ps,
+    failed = failed orelse exists (fn SOME (Exn.Exn _) => true | _ => false) ps}
+  end;
+
 
 (* fulfilled proofs *)
 
-fun raw_proof_of (Thm (Deriv {body, ...}, _)) = Proofterm.proof_of body;
-
 fun proof_body_of (Thm (Deriv {open_promises, promises, body, ...}, {thy_ref, ...})) =
   let
     val _ = Exn.release_all (map (Future.join_result o #2) (rev open_promises));
-    val ps = map (apsnd (raw_proof_of o Future.join)) promises;
+    val ps = map (apsnd (raw_proof_body_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;
@@ -1666,7 +1680,7 @@
     val {thy_ref, hyps, prop, tpairs, ...} = args;
     val _ = null tpairs orelse raise THM ("put_name: unsolved flex-flex constraints", 0, [thm]);
 
-    val ps = map (apsnd (Future.map proof_of)) promises;
+    val ps = map (apsnd (Future.map proof_body_of)) promises;
     val thy = Theory.deref thy_ref;
     val (pthm, proof) = Pt.thm_proof thy name hyps prop ps body;
 
@@ -1687,8 +1701,8 @@
     if T <> propT then
       raise THM ("Oracle's result must have type prop: " ^ name, 0, [])
     else
-      let val prf = Pt.oracle_proof name prop in
-        Thm (make_deriv ~1 [] [] (Pt.make_oracles prf) [] prf,
+      let val (ora, prf) = Pt.oracle_proof name prop in
+        Thm (make_deriv ~1 [] [] [ora] [] prf,
          {thy_ref = Theory.merge_refs (thy_ref1, thy_ref2),
           tags = [],
           maxidx = maxidx,
--- a/src/Sequents/LK.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Sequents/LK.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -226,9 +226,9 @@
 
 lemma split_if: "|- P(if Q then x else y) <-> ((Q --> P(x)) & (~Q --> P(y)))"
   apply (rule_tac P = Q in cut)
-   apply (tactic {* simp_tac (simpset () addsimps @{thms if_P}) 2 *})
+   apply (tactic {* simp_tac (@{simpset} addsimps @{thms if_P}) 2 *})
   apply (rule_tac P = "~Q" in cut)
-   apply (tactic {* simp_tac (simpset() addsimps @{thms if_not_P}) 2 *})
+   apply (tactic {* simp_tac (@{simpset} addsimps @{thms if_not_P}) 2 *})
   apply (tactic {* fast_tac LK_pack 1 *})
   done
 
--- a/src/Tools/Compute_Oracle/am_compiler.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Tools/Compute_Oracle/am_compiler.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -185,7 +185,7 @@
 
     in
 	compiled_rewriter := NONE;	
-	use_text ML_Context.name_space (1, "") Output.ml_output false (!buffer);
+	use_text ML_Context.local_context (1, "") false (!buffer);
 	case !compiled_rewriter of 
 	    NONE => raise (Compile "cannot communicate with compiled function")
 	  | SOME r => (compiled_rewriter := NONE; r)
--- a/src/Tools/Compute_Oracle/am_sml.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Tools/Compute_Oracle/am_sml.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -492,7 +492,7 @@
 
 fun writeTextFile name s = File.write (Path.explode name) s
 
-fun use_source src = use_text ML_Context.name_space (1, "") Output.ml_output false src
+fun use_source src = use_text ML_Context.local_context (1, "") false src
     
 fun compile cache_patterns const_arity eqs = 
     let
--- a/src/Tools/code/code_haskell.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Tools/code/code_haskell.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -42,18 +42,18 @@
      of [] => []
       | classbinds => Pretty.enum "," "(" ")" (
           map (fn (v, class) =>
-            str (class_name class ^ " " ^ Code_Name.lookup_var tyvars v)) classbinds)
+            str (class_name class ^ " " ^ Code_Printer.lookup_var tyvars v)) classbinds)
           @@ str " => ";
     fun pr_typforall tyvars vs = case map fst vs
      of [] => []
       | vnames => str "forall " :: Pretty.breaks
-          (map (str o Code_Name.lookup_var tyvars) vnames) @ str "." @@ Pretty.brk 1;
+          (map (str o Code_Printer.lookup_var tyvars) vnames) @ str "." @@ Pretty.brk 1;
     fun pr_tycoexpr tyvars fxy (tyco, tys) =
       brackify fxy (str tyco :: map (pr_typ tyvars BR) tys)
     and pr_typ tyvars fxy (tycoexpr as tyco `%% tys) = (case syntax_tyco tyco
          of NONE => pr_tycoexpr tyvars fxy (deresolve tyco, tys)
           | SOME (i, pr) => pr (pr_typ tyvars) fxy tys)
-      | pr_typ tyvars fxy (ITyVar v) = (str o Code_Name.lookup_var tyvars) v;
+      | pr_typ tyvars fxy (ITyVar v) = (str o Code_Printer.lookup_var tyvars) v;
     fun pr_typdecl tyvars (vs, tycoexpr) =
       Pretty.block (pr_typcontext tyvars vs @| pr_tycoexpr tyvars NOBR tycoexpr);
     fun pr_typscheme tyvars (vs, ty) =
@@ -69,7 +69,7 @@
                   pr_term tyvars thm vars BR t2
                 ])
       | pr_term tyvars thm vars fxy (IVar v) =
-          (str o Code_Name.lookup_var vars) v
+          (str o Code_Printer.lookup_var vars) v
       | pr_term tyvars thm vars fxy (t as _ `|-> _) =
           let
             val (binds, t') = Code_Thingol.unfold_abs t;
@@ -127,7 +127,7 @@
       | pr_case tyvars thm vars fxy ((_, []), _) = str "error \"empty case\"";
     fun pr_stmt (name, Code_Thingol.Fun (_, ((vs, ty), []))) =
           let
-            val tyvars = Code_Name.intro_vars (map fst vs) init_syms;
+            val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
             val n = (length o fst o Code_Thingol.unfold_fun) ty;
           in
             Pretty.chunks [
@@ -150,7 +150,7 @@
       | pr_stmt (name, Code_Thingol.Fun (_, ((vs, ty), raw_eqs))) =
           let
             val eqs = filter (snd o snd) raw_eqs;
-            val tyvars = Code_Name.intro_vars (map fst vs) init_syms;
+            val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
             fun pr_eq ((ts, t), (thm, _)) =
               let
                 val consts = map_filter
@@ -158,8 +158,8 @@
                     then NONE else (SOME o Long_Name.base_name o deresolve) c)
                     ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []);
                 val vars = init_syms
-                  |> Code_Name.intro_vars consts
-                  |> Code_Name.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
+                  |> Code_Printer.intro_vars consts
+                  |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
                        (insert (op =)) ts []);
               in
                 semicolon (
@@ -182,7 +182,7 @@
           end
       | pr_stmt (name, Code_Thingol.Datatype (_, (vs, []))) =
           let
-            val tyvars = Code_Name.intro_vars (map fst vs) init_syms;
+            val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
           in
             semicolon [
               str "data",
@@ -191,7 +191,7 @@
           end
       | pr_stmt (name, Code_Thingol.Datatype (_, (vs, [(co, [ty])]))) =
           let
-            val tyvars = Code_Name.intro_vars (map fst vs) init_syms;
+            val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
           in
             semicolon (
               str "newtype"
@@ -204,7 +204,7 @@
           end
       | pr_stmt (name, Code_Thingol.Datatype (_, (vs, co :: cos))) =
           let
-            val tyvars = Code_Name.intro_vars (map fst vs) init_syms;
+            val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
             fun pr_co (co, tys) =
               concat (
                 (str o deresolve_base) co
@@ -222,7 +222,7 @@
           end
       | pr_stmt (name, Code_Thingol.Class (_, (v, (superclasses, classparams)))) =
           let
-            val tyvars = Code_Name.intro_vars [v] init_syms;
+            val tyvars = Code_Printer.intro_vars [v] init_syms;
             fun pr_classparam (classparam, ty) =
               semicolon [
                 (str o deresolve_base) classparam,
@@ -234,7 +234,7 @@
               Pretty.block [
                 str "class ",
                 Pretty.block (pr_typcontext tyvars [(v, map fst superclasses)]),
-                str (deresolve_base name ^ " " ^ Code_Name.lookup_var tyvars v),
+                str (deresolve_base name ^ " " ^ Code_Printer.lookup_var tyvars v),
                 str " where {"
               ],
               str "};"
@@ -244,7 +244,7 @@
           let
             val split_abs_pure = (fn (v, _) `|-> t => SOME (v, t) | _ => NONE);
             val unfold_abs_pure = Code_Thingol.unfoldr split_abs_pure;
-            val tyvars = Code_Name.intro_vars (map fst vs) init_syms;
+            val tyvars = Code_Printer.intro_vars (map fst vs) init_syms;
             fun pr_instdef ((classparam, c_inst), (thm, _)) = case syntax_const classparam
              of NONE => semicolon [
                     (str o deresolve_base) classparam,
@@ -259,8 +259,8 @@
                     val proto_rhs = Code_Thingol.eta_expand k (c_inst, []);
                     val (vs, rhs) = unfold_abs_pure proto_rhs;
                     val vars = init_syms
-                      |> Code_Name.intro_vars (the_list const)
-                      |> Code_Name.intro_vars vs;
+                      |> Code_Printer.intro_vars (the_list const)
+                      |> Code_Printer.intro_vars vs;
                     val lhs = IConst (classparam, ([], tys)) `$$ map IVar vs;
                       (*dictionaries are not relevant at this late stage*)
                   in
@@ -288,20 +288,20 @@
   let
     val module_alias = if is_some module_name then K module_name else raw_module_alias;
     val reserved_names = Name.make_context reserved_names;
-    val mk_name_module = Code_Name.mk_name_module reserved_names module_prefix module_alias program;
+    val mk_name_module = Code_Printer.mk_name_module reserved_names module_prefix module_alias program;
     fun add_stmt (name, (stmt, deps)) =
       let
-        val (module_name, base) = Code_Name.dest_name name;
+        val (module_name, base) = Code_Printer.dest_name name;
         val module_name' = mk_name_module module_name;
         val mk_name_stmt = yield_singleton Name.variants;
         fun add_fun upper (nsp_fun, nsp_typ) =
           let
             val (base', nsp_fun') =
-              mk_name_stmt (if upper then Code_Name.first_upper base else base) nsp_fun
+              mk_name_stmt (if upper then Code_Printer.first_upper base else base) nsp_fun
           in (base', (nsp_fun', nsp_typ)) end;
         fun add_typ (nsp_fun, nsp_typ) =
           let
-            val (base', nsp_typ') = mk_name_stmt (Code_Name.first_upper base) nsp_typ
+            val (base', nsp_typ') = mk_name_stmt (Code_Printer.first_upper base) nsp_typ
           in (base', (nsp_fun, nsp_typ')) end;
         val add_name = case stmt
          of Code_Thingol.Fun _ => add_fun false
@@ -330,7 +330,7 @@
       (Graph.get_node program name, Graph.imm_succs program name))
       (Graph.strong_conn program |> flat)) Symtab.empty;
     fun deresolver name = (fst o the o AList.lookup (op =) ((fst o snd o the
-      o Symtab.lookup hs_program) ((mk_name_module o fst o Code_Name.dest_name) name))) name
+      o Symtab.lookup hs_program) ((mk_name_module o fst o Code_Printer.dest_name) name))) name
       handle Option => error ("Unknown statement name: " ^ labelled_name name);
   in (deresolver, hs_program) end;
 
@@ -357,7 +357,7 @@
               andalso forall (deriv' tycos) tys
           | deriv' _ (ITyVar _) = true
       in deriv [] tyco end;
-    val reserved_names = Code_Name.make_vars reserved_names;
+    val reserved_names = Code_Printer.make_vars reserved_names;
     fun pr_stmt qualified = pr_haskell_stmt naming labelled_name
       syntax_class syntax_tyco syntax_const reserved_names
       (if qualified then deresolver else Long_Name.base_name o deresolver)
--- a/src/Tools/code/code_ml.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Tools/code/code_ml.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -50,8 +50,8 @@
     val pr_label_classparam = Long_Name.base_name o Long_Name.qualifier;
     fun pr_dicts fxy ds =
       let
-        fun pr_dictvar (v, (_, 1)) = Code_Name.first_upper v ^ "_"
-          | pr_dictvar (v, (i, _)) = Code_Name.first_upper v ^ string_of_int (i+1) ^ "_";
+        fun pr_dictvar (v, (_, 1)) = Code_Printer.first_upper v ^ "_"
+          | pr_dictvar (v, (i, _)) = Code_Printer.first_upper v ^ string_of_int (i+1) ^ "_";
         fun pr_proj [] p =
               p
           | pr_proj [p'] p =
@@ -87,7 +87,7 @@
     fun pr_term is_closure thm vars fxy (IConst c) =
           pr_app is_closure thm vars fxy (c, [])
       | pr_term is_closure thm vars fxy (IVar v) =
-          str (Code_Name.lookup_var vars v)
+          str (Code_Printer.lookup_var vars v)
       | pr_term is_closure thm vars fxy (t as t1 `$ t2) =
           (case Code_Thingol.unfold_const_app t
            of SOME c_ts => pr_app is_closure thm vars fxy c_ts
@@ -182,7 +182,7 @@
                 then NONE else (SOME o Long_Name.base_name o deresolve) c)
                 (Code_Thingol.fold_constnames (insert (op =)) t []);
             val vars = reserved_names
-              |> Code_Name.intro_vars consts;
+              |> Code_Printer.intro_vars consts;
           in
             concat [
               str "val",
@@ -207,8 +207,8 @@
                         then NONE else (SOME o Long_Name.base_name o deresolve) c)
                         ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []);
                     val vars = reserved_names
-                      |> Code_Name.intro_vars consts
-                      |> Code_Name.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
+                      |> Code_Printer.intro_vars consts
+                      |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
                            (insert (op =)) ts []);
                   in
                     concat (
@@ -266,7 +266,7 @@
           in Pretty.chunks (ps @| Pretty.block ([p, str ";"])) end
      | pr_stmt (MLClass (class, (v, (superclasses, classparams)))) =
           let
-            val w = Code_Name.first_upper v ^ "_";
+            val w = Code_Printer.first_upper v ^ "_";
             fun pr_superclass_field (class, classrel) =
               (concat o map str) [
                 pr_label_classrel classrel, ":", "'" ^ v, deresolve class
@@ -362,8 +362,8 @@
   let
     fun pr_dicts fxy ds =
       let
-        fun pr_dictvar (v, (_, 1)) = "_" ^ Code_Name.first_upper v
-          | pr_dictvar (v, (i, _)) = "_" ^ Code_Name.first_upper v ^ string_of_int (i+1);
+        fun pr_dictvar (v, (_, 1)) = "_" ^ Code_Printer.first_upper v
+          | pr_dictvar (v, (i, _)) = "_" ^ Code_Printer.first_upper v ^ string_of_int (i+1);
         fun pr_proj ps p =
           fold_rev (fn p2 => fn p1 => Pretty.block [p1, str ".", str p2]) ps p
         fun pr_dict fxy (DictConst (inst, dss)) =
@@ -395,7 +395,7 @@
     fun pr_term is_closure thm vars fxy (IConst c) =
           pr_app is_closure thm vars fxy (c, [])
       | pr_term is_closure thm vars fxy (IVar v) =
-          str (Code_Name.lookup_var vars v)
+          str (Code_Printer.lookup_var vars v)
       | pr_term is_closure thm vars fxy (t as t1 `$ t2) =
           (case Code_Thingol.unfold_const_app t
            of SOME c_ts => pr_app is_closure thm vars fxy c_ts
@@ -468,8 +468,8 @@
         val x = Name.variant (map_filter I fished1) "x";
         val fished2 = map_index (fillup_param x) fished1;
         val (fished3, _) = Name.variants fished2 Name.context;
-        val vars' = Code_Name.intro_vars fished3 vars;
-      in map (Code_Name.lookup_var vars') fished3 end;
+        val vars' = Code_Printer.intro_vars fished3 vars;
+      in map (Code_Printer.lookup_var vars') fished3 end;
     fun pr_stmt (MLExc (name, n)) =
           let
             val exc_str =
@@ -491,7 +491,7 @@
                 then NONE else (SOME o Long_Name.base_name o deresolve) c)
                 (Code_Thingol.fold_constnames (insert (op =)) t []);
             val vars = reserved_names
-              |> Code_Name.intro_vars consts;
+              |> Code_Printer.intro_vars consts;
           in
             concat [
               str "let",
@@ -511,8 +511,8 @@
                     then NONE else (SOME o Long_Name.base_name o deresolve) c)
                     ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []);
                 val vars = reserved_names
-                  |> Code_Name.intro_vars consts
-                  |> Code_Name.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
+                  |> Code_Printer.intro_vars consts
+                  |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
                       (insert (op =)) ts []);
               in concat [
                 (Pretty.block o Pretty.commas)
@@ -527,8 +527,8 @@
                         then NONE else (SOME o Long_Name.base_name o deresolve) c)
                         ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []);
                     val vars = reserved_names
-                      |> Code_Name.intro_vars consts
-                      |> Code_Name.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
+                      |> Code_Printer.intro_vars consts
+                      |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames)
                           (insert (op =)) ts []);
                   in
                     concat (
@@ -556,7 +556,7 @@
                         ((fold o Code_Thingol.fold_constnames)
                           (insert (op =)) (map (snd o fst) eqs) []);
                     val vars = reserved_names
-                      |> Code_Name.intro_vars consts;
+                      |> Code_Printer.intro_vars consts;
                     val dummy_parms = (map str o fish_params vars o map (fst o fst)) eqs;
                   in
                     Pretty.block (
@@ -623,7 +623,7 @@
           in Pretty.chunks (ps @| Pretty.block ([p, str ";;"])) end
      | pr_stmt (MLClass (class, (v, (superclasses, classparams)))) =
           let
-            val w = "_" ^ Code_Name.first_upper v;
+            val w = "_" ^ Code_Printer.first_upper v;
             fun pr_superclass_field (class, classrel) =
               (concat o map str) [
                 deresolve classrel, ":", "'" ^ v, deresolve class
@@ -765,11 +765,11 @@
       let
         val (x, nsp_typ') = f nsp_typ
       in (x, (nsp_fun, nsp_typ')) end;
-    val mk_name_module = Code_Name.mk_name_module reserved_names NONE module_alias program;
+    val mk_name_module = Code_Printer.mk_name_module reserved_names NONE module_alias program;
     fun mk_name_stmt upper name nsp =
       let
-        val (_, base) = Code_Name.dest_name name;
-        val base' = if upper then Code_Name.first_upper base else base;
+        val (_, base) = Code_Printer.dest_name name;
+        val base' = if upper then Code_Printer.first_upper base else base;
         val ([base''], nsp') = Name.variants [base'] nsp;
       in (base'', nsp') end;
     fun rearrange_fun name (tysm as (vs, ty), raw_eqs) =
@@ -853,7 +853,7 @@
           []
           |> fold (fold (insert (op =)) o Graph.imm_succs program) names
           |> subtract (op =) names;
-        val (module_names, _) = (split_list o map Code_Name.dest_name) names;
+        val (module_names, _) = (split_list o map Code_Printer.dest_name) names;
         val module_name = (the_single o distinct (op =) o map mk_name_module) module_names
           handle Empty =>
             error ("Different namespace prefixes for mutual dependencies:\n"
@@ -863,7 +863,7 @@
         val module_name_path = Long_Name.explode module_name;
         fun add_dep name name' =
           let
-            val module_name' = (mk_name_module o fst o Code_Name.dest_name) name';
+            val module_name' = (mk_name_module o fst o Code_Printer.dest_name) name';
           in if module_name = module_name' then
             map_node module_name_path (Graph.add_edge (name, name'))
           else let
@@ -891,7 +891,7 @@
           (rev (Graph.strong_conn program)));
     fun deresolver prefix name = 
       let
-        val module_name = (fst o Code_Name.dest_name) name;
+        val module_name = (fst o Code_Printer.dest_name) name;
         val module_name' = (Long_Name.explode o mk_name_module) module_name;
         val (_, (_, remainder)) = chop_prefix (op =) (prefix, module_name');
         val stmt_name =
@@ -914,7 +914,7 @@
     val module_name = if null stmt_names then raw_module_name else SOME "Code";
     val (deresolver, nodes) = ml_node_of_program labelled_name module_name
       reserved_names raw_module_alias program;
-    val reserved_names = Code_Name.make_vars reserved_names;
+    val reserved_names = Code_Printer.make_vars reserved_names;
     fun pr_node prefix (Dummy _) =
           NONE
       | pr_node prefix (Stmt (_, stmt)) = if null stmt_names orelse
@@ -969,7 +969,7 @@
         val (value_code, [SOME value_name']) = ml_code_of thy naming program' [value_name];
         val sml_code = "let\n" ^ value_code ^ "\nin " ^ value_name'
           ^ space_implode " " (map (enclose "(" ")") args) ^ " end";
-      in ML_Context.evaluate ctxt Output.ml_output false reff sml_code end;
+      in ML_Context.evaluate ctxt false reff sml_code end;
   in eval'' thy (rpair eval') ct end;
 
 fun eval_term reff = eval Code_Thingol.eval_term I reff;
@@ -1037,7 +1037,7 @@
 fun isar_seri_sml module_name =
   Code_Target.parse_args (Scan.succeed ())
   #> (fn () => serialize_ml target_SML
-      (SOME (use_text ML_Context.name_space (1, "generated code") Output.ml_output false))
+      (SOME (use_text ML_Context.local_context (1, "generated code") false))
       pr_sml_module pr_sml_stmt module_name);
 
 fun isar_seri_ocaml module_name =
--- a/src/Tools/code/code_name.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Tools/code/code_name.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -6,45 +6,20 @@
 
 signature CODE_NAME =
 sig
-  structure StringPairTab: TABLE
-  val first_upper: string -> string
-  val first_lower: string -> string
-  val dest_name: string -> string * string
-
   val purify_var: string -> string
   val purify_tvar: string -> string
-  val purify_sym: string -> string
-  val purify_base: bool -> string -> string
+  val purify_base: string -> string
   val check_modulename: string -> string
 
-  type var_ctxt
-  val make_vars: string list -> var_ctxt
-  val intro_vars: string list -> var_ctxt -> var_ctxt
-  val lookup_var: var_ctxt -> string -> string
-
   val read_const_exprs: theory -> string list -> string list * string list
-  val mk_name_module: Name.context -> string option -> (string -> string option)
-    -> 'a Graph.T -> string -> string
 end;
 
 structure Code_Name: CODE_NAME =
 struct
 
-(** auxiliary **)
-
-structure StringPairTab =
-  TableFun(type key = string * string val ord = prod_ord fast_string_ord fast_string_ord);
-
-val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode;
-val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o explode;
-
-val dest_name =
-  apfst Long_Name.implode o split_last o fst o split_last o Long_Name.explode;
-
-
 (** purification **)
 
-fun purify_name upper lower =
+fun purify_name upper =
   let
     fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s orelse s = "'";
     val is_junk = not o is_valid andf Symbol.is_regular;
@@ -55,9 +30,8 @@
         --| junk))
       ::: Scan.repeat ((Scan.many1 is_valid >> implode) --| junk));
     fun upper_lower cs = if upper then nth_map 0 Symbol.to_ascii_upper cs
-      else if lower then (if forall Symbol.is_ascii_upper cs
-        then map else nth_map 0) Symbol.to_ascii_lower cs
-      else cs;
+      else (if forall Symbol.is_ascii_upper cs
+        then map else nth_map 0) Symbol.to_ascii_lower cs;
   in
     explode
     #> scan_valids
@@ -68,7 +42,7 @@
   end;
 
 fun purify_var "" = "x"
-  | purify_var v = purify_name false true v;
+  | purify_var v = purify_name false v;
 
 fun purify_tvar "" = "'a"
   | purify_tvar v =
@@ -81,52 +55,29 @@
       (Scan.repeat ($$ "_" |-- $$ "_" >> (fn _ => ".") || Scan.one Symbol.is_regular))
   #> implode
   #> Long_Name.explode
-  #> map (purify_name true false);
+  #> map (purify_name true);
 
 (*FIMXE non-canonical function treating non-canonical names*)
-fun purify_base _ "op &" = "and"
-  | purify_base _ "op |" = "or"
-  | purify_base _ "op -->" = "implies"
-  | purify_base _ "{}" = "empty"
-  | purify_base _ "op :" = "member"
-  | purify_base _ "op Int" = "intersect"
-  | purify_base _ "op Un" = "union"
-  | purify_base _ "*" = "product"
-  | purify_base _ "+" = "sum"
-  | purify_base lower s = if String.isPrefix "op =" s
-      then "eq" ^ purify_name false lower s
-      else purify_name false lower s;
-
-val purify_sym = purify_base false;
+fun purify_base "op &" = "and"
+  | purify_base "op |" = "or"
+  | purify_base "op -->" = "implies"
+  | purify_base "op :" = "member"
+  | purify_base "*" = "product"
+  | purify_base "+" = "sum"
+  | purify_base s = if String.isPrefix "op =" s
+      then "eq" ^ purify_name false s
+      else purify_name false s;
 
 fun check_modulename mn =
   let
     val mns = Long_Name.explode mn;
-    val mns' = map (purify_name true false) mns;
+    val mns' = map (purify_name true) mns;
   in
     if mns' = mns then mn else error ("Invalid module name: " ^ quote mn ^ "\n"
       ^ "perhaps try " ^ quote (Long_Name.implode mns'))
   end;
 
 
-(** variable name contexts **)
-
-type var_ctxt = string Symtab.table * Name.context;
-
-fun make_vars names = (fold (fn name => Symtab.update_new (name, name)) names Symtab.empty,
-  Name.make_context names);
-
-fun intro_vars names (namemap, namectxt) =
-  let
-    val (names', namectxt') = Name.variants names namectxt;
-    val namemap' = fold2 (curry Symtab.update) names names' namemap;
-  in (namemap', namectxt') end;
-
-fun lookup_var (namemap, _) name = case Symtab.lookup namemap name
- of SOME name' => name'
-  | NONE => error ("Invalid name in context: " ^ quote name);
-
-
 (** misc **)
 
 fun read_const_exprs thy =
@@ -150,22 +101,4 @@
           else ([Code_Unit.read_const thy s], []);
   in pairself flat o split_list o map read_const_expr end;
 
-fun mk_name_module reserved_names module_prefix module_alias program =
-  let
-    fun mk_alias name = case module_alias name
-     of SOME name' => name'
-      | NONE => name
-          |> Long_Name.explode
-          |> map (fn name => (the_single o fst) (Name.variants [name] reserved_names))
-          |> Long_Name.implode;
-    fun mk_prefix name = case module_prefix
-     of SOME module_prefix => Long_Name.append module_prefix name
-      | NONE => name;
-    val tab =
-      Symtab.empty
-      |> Graph.fold ((fn name => Symtab.default (name, (mk_alias #> mk_prefix) name))
-           o fst o dest_name o fst)
-             program
-  in the o Symtab.lookup tab end;
-
 end;
--- a/src/Tools/code/code_printer.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Tools/code/code_printer.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -16,6 +16,13 @@
   val semicolon: Pretty.T list -> Pretty.T
   val enum_default: string -> string -> string -> string -> Pretty.T list -> Pretty.T
 
+  val first_upper: string -> string
+  val first_lower: string -> string
+  type var_ctxt
+  val make_vars: string list -> var_ctxt
+  val intro_vars: string list -> var_ctxt -> var_ctxt
+  val lookup_var: var_ctxt -> string -> string
+
   type lrx
   val L: lrx
   val R: lrx
@@ -42,14 +49,14 @@
     -> fixity -> 'a list -> Pretty.T)) option * OuterParse.token list
   val simple_const_syntax: (int * ((fixity -> iterm -> Pretty.T)
     -> fixity -> (iterm * itype) list -> Pretty.T)) option -> const_syntax option
-  val gen_pr_app: (thm -> Code_Name.var_ctxt -> const * iterm list -> Pretty.T list)
-    -> (thm -> Code_Name.var_ctxt -> fixity -> iterm -> Pretty.T)
+  val gen_pr_app: (thm -> var_ctxt -> const * iterm list -> Pretty.T list)
+    -> (thm -> var_ctxt -> fixity -> iterm -> Pretty.T)
     -> (string -> const_syntax option) -> Code_Thingol.naming
-    -> thm -> Code_Name.var_ctxt -> fixity -> const * iterm list -> Pretty.T
+    -> thm -> var_ctxt -> fixity -> const * iterm list -> Pretty.T
   val gen_pr_bind: ((string option * Pretty.T option) * itype -> Pretty.T)
-    -> (thm -> Code_Name.var_ctxt -> fixity -> iterm -> Pretty.T)
+    -> (thm -> var_ctxt -> fixity -> iterm -> Pretty.T)
     -> thm -> fixity
-    -> (string option * iterm option) * itype -> Code_Name.var_ctxt -> Pretty.T * Code_Name.var_ctxt
+    -> (string option * iterm option) * itype -> var_ctxt -> Pretty.T * var_ctxt
 
   type literals
   val Literals: { literal_char: string -> string, literal_string: string -> string,
@@ -60,6 +67,10 @@
   val literal_numeral: literals -> bool -> int -> string
   val literal_list: literals -> Pretty.T list -> Pretty.T
   val infix_cons: literals -> int * string
+
+  val mk_name_module: Name.context -> string option -> (string -> string option)
+    -> 'a Graph.T -> string -> string
+  val dest_name: string -> string * string
 end;
 
 structure Code_Printer : CODE_PRINTER =
@@ -83,6 +94,27 @@
   | enum_default default sep opn cls xs = Pretty.enum sep opn cls xs;
 
 
+(** names and variable name contexts **)
+
+type var_ctxt = string Symtab.table * Name.context;
+
+fun make_vars names = (fold (fn name => Symtab.update_new (name, name)) names Symtab.empty,
+  Name.make_context names);
+
+fun intro_vars names (namemap, namectxt) =
+  let
+    val (names', namectxt') = Name.variants names namectxt;
+    val namemap' = fold2 (curry Symtab.update) names names' namemap;
+  in (namemap', namectxt') end;
+
+fun lookup_var (namemap, _) name = case Symtab.lookup namemap name
+ of SOME name' => name'
+  | NONE => error ("Invalid name in context: " ^ quote name);
+
+val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode;
+val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o explode;
+
+
 (** syntax printer **)
 
 (* binding priorities *)
@@ -125,8 +157,8 @@
 
 type tyco_syntax = int * ((fixity -> itype -> Pretty.T)
   -> fixity -> itype list -> Pretty.T);
-type const_syntax = int * ((Code_Name.var_ctxt -> fixity -> iterm -> Pretty.T)
-  -> Code_Thingol.naming -> thm -> Code_Name.var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T);
+type const_syntax = int * ((var_ctxt -> fixity -> iterm -> Pretty.T)
+  -> Code_Thingol.naming -> thm -> var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T);
 
 fun simple_const_syntax x = (Option.map o apsnd)
   (fn pretty => fn pr => fn naming => fn thm => fn vars => pretty (pr vars)) x;
@@ -150,9 +182,9 @@
     val vs = case pat
      of SOME pat => Code_Thingol.fold_varnames (insert (op =)) pat []
       | NONE => [];
-    val vars' = Code_Name.intro_vars (the_list v) vars;
-    val vars'' = Code_Name.intro_vars vs vars';
-    val v' = Option.map (Code_Name.lookup_var vars') v;
+    val vars' = intro_vars (the_list v) vars;
+    val vars'' = intro_vars vs vars';
+    val v' = Option.map (lookup_var vars') v;
     val pat' = Option.map (pr_term thm vars'' fxy) pat;
   in (pr_bind ((v', pat'), ty), vars'') end;
 
@@ -239,4 +271,28 @@
 val literal_list = #literal_list o dest_Literals;
 val infix_cons = #infix_cons o dest_Literals;
 
+
+(** module name spaces **)
+
+val dest_name =
+  apfst Long_Name.implode o split_last o fst o split_last o Long_Name.explode;
+
+fun mk_name_module reserved_names module_prefix module_alias program =
+  let
+    fun mk_alias name = case module_alias name
+     of SOME name' => name'
+      | NONE => name
+          |> Long_Name.explode
+          |> map (fn name => (the_single o fst) (Name.variants [name] reserved_names))
+          |> Long_Name.implode;
+    fun mk_prefix name = case module_prefix
+     of SOME module_prefix => Long_Name.append module_prefix name
+      | NONE => name;
+    val tab =
+      Symtab.empty
+      |> Graph.fold ((fn name => Symtab.default (name, (mk_alias #> mk_prefix) name))
+           o fst o dest_name o fst)
+             program
+  in the o Symtab.lookup tab end;
+
 end; (*struct*)
--- a/src/Tools/code/code_target.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Tools/code/code_target.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -82,11 +82,9 @@
 
 (** theory data **)
 
-structure StringPairTab = Code_Name.StringPairTab;
-
 datatype name_syntax_table = NameSyntaxTable of {
   class: string Symtab.table,
-  instance: unit StringPairTab.table,
+  instance: unit Symreltab.table,
   tyco: tyco_syntax Symtab.table,
   const: const_syntax Symtab.table
 };
@@ -99,7 +97,7 @@
     NameSyntaxTable { class = class2, instance = instance2, tyco = tyco2, const = const2 }) =
   mk_name_syntax_table (
     (Symtab.join (K snd) (class1, class2),
-       StringPairTab.join (K snd) (instance1, instance2)),
+       Symreltab.join (K snd) (instance1, instance2)),
     (Symtab.join (K snd) (tyco1, tyco2),
        Symtab.join (K snd) (const1, const2))
   );
@@ -194,7 +192,7 @@
     thy
     |> (CodeTargetData.map o apfst oo Symtab.map_default)
           (target, mk_target ((serial (), seri), (([], Symtab.empty),
-            (mk_name_syntax_table ((Symtab.empty, StringPairTab.empty), (Symtab.empty, Symtab.empty)),
+            (mk_name_syntax_table ((Symtab.empty, Symreltab.empty), (Symtab.empty, Symtab.empty)),
               Symtab.empty))))
           ((map_target o apfst o apsnd o K) seri)
   end;
@@ -262,11 +260,11 @@
   in if add_del then
     thy
     |> (map_name_syntax target o apfst o apsnd)
-        (StringPairTab.update (inst, ()))
+        (Symreltab.update (inst, ()))
   else
     thy
     |> (map_name_syntax target o apfst o apsnd)
-        (StringPairTab.delete_safe inst)
+        (Symreltab.delete_safe inst)
   end;
 
 fun gen_add_syntax_tyco prep_tyco target raw_tyco raw_syn thy =
@@ -407,7 +405,7 @@
       |>> map_filter I;
     val (names_class, class') = distill_names Code_Thingol.lookup_class class;
     val names_inst = map_filter (Code_Thingol.lookup_instance naming)
-      (StringPairTab.keys instance);
+      (Symreltab.keys instance);
     val (names_tyco, tyco') = distill_names Code_Thingol.lookup_tyco tyco;
     val (names_const, const') = distill_names Code_Thingol.lookup_const const;
     val names_hidden = names_class @ names_inst @ names_tyco @ names_const;
--- a/src/Tools/code/code_thingol.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Tools/code/code_thingol.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -242,7 +242,7 @@
   fun namify thy get_basename get_thyname name =
     let
       val prefix = get_thyname thy name;
-      val base = (Code_Name.purify_base true o get_basename) name;
+      val base = (Code_Name.purify_base o get_basename) name;
     in Long_Name.append prefix base end;
 in
 
@@ -261,13 +261,11 @@
 
 (* data *)
 
-structure StringPairTab = Code_Name.StringPairTab;
-
 datatype naming = Naming of {
   class: class Symtab.table * Name.context,
-  classrel: string StringPairTab.table * Name.context,
+  classrel: string Symreltab.table * Name.context,
   tyco: string Symtab.table * Name.context,
-  instance: string StringPairTab.table * Name.context,
+  instance: string Symreltab.table * Name.context,
   const: string Symtab.table * Name.context
 }
 
@@ -275,9 +273,9 @@
 
 val empty_naming = Naming {
   class = (Symtab.empty, Name.context),
-  classrel = (StringPairTab.empty, Name.context),
+  classrel = (Symreltab.empty, Name.context),
   tyco = (Symtab.empty, Name.context),
-  instance = (StringPairTab.empty, Name.context),
+  instance = (Symreltab.empty, Name.context),
   const = (Symtab.empty, Name.context)
 };
 
@@ -334,22 +332,22 @@
 val lookup_class = add_suffix suffix_class
   oo Symtab.lookup o fst o #class o dest_Naming;
 val lookup_classrel = add_suffix suffix_classrel
-  oo StringPairTab.lookup o fst o #classrel o dest_Naming;
+  oo Symreltab.lookup o fst o #classrel o dest_Naming;
 val lookup_tyco = add_suffix suffix_tyco
   oo Symtab.lookup o fst o #tyco o dest_Naming;
 val lookup_instance = add_suffix suffix_instance
-  oo StringPairTab.lookup o fst o #instance o dest_Naming;
+  oo Symreltab.lookup o fst o #instance o dest_Naming;
 val lookup_const = add_suffix suffix_const
   oo Symtab.lookup o fst o #const o dest_Naming;
 
 fun declare_class thy = declare thy map_class
   lookup_class Symtab.update_new namify_class;
 fun declare_classrel thy = declare thy map_classrel
-  lookup_classrel StringPairTab.update_new namify_classrel;
+  lookup_classrel Symreltab.update_new namify_classrel;
 fun declare_tyco thy = declare thy map_tyco
   lookup_tyco Symtab.update_new namify_tyco;
 fun declare_instance thy = declare thy map_instance
-  lookup_instance StringPairTab.update_new namify_instance;
+  lookup_instance Symreltab.update_new namify_instance;
 fun declare_const thy = declare thy map_const
   lookup_const Symtab.update_new namify_const;
 
--- a/src/Tools/induct.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Tools/induct.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -259,16 +259,15 @@
   spec setN Args.const >> add_pred ||
   Scan.lift Args.del >> K del;
 
-val cases_att = attrib cases_type cases_pred cases_del;
-val induct_att = attrib induct_type induct_pred induct_del;
-val coinduct_att = attrib coinduct_type coinduct_pred coinduct_del;
-
 in
 
 val attrib_setup =
-  Attrib.setup @{binding cases} cases_att "declaration of cases rule" #>
-  Attrib.setup @{binding induct} induct_att "declaration of induction rule" #>
-  Attrib.setup @{binding coinduct} coinduct_att "declaration of coinduction rule";
+  Attrib.setup @{binding cases} (attrib cases_type cases_pred cases_del)
+    "declaration of cases rule" #>
+  Attrib.setup @{binding induct} (attrib induct_type induct_pred induct_del)
+    "declaration of induction rule" #>
+  Attrib.setup @{binding coinduct} (attrib coinduct_type coinduct_pred coinduct_del)
+    "declaration of coinduction rule";
 
 end;
 
@@ -735,23 +734,29 @@
 
 in
 
-val cases_meth =
-  P.and_list' (Scan.repeat (unless_more_args inst)) -- Scan.option cases_rule >>
-  (fn (insts, opt_rule) => fn ctxt =>
-    METHOD_CASES (fn facts => Seq.DETERM (HEADGOAL (cases_tac ctxt insts opt_rule facts))));
+val cases_setup =
+  Method.setup @{binding cases}
+    (P.and_list' (Scan.repeat (unless_more_args inst)) -- Scan.option cases_rule >>
+      (fn (insts, opt_rule) => fn ctxt =>
+        METHOD_CASES (fn facts => Seq.DETERM (HEADGOAL (cases_tac ctxt insts opt_rule facts)))))
+    "case analysis on types or predicates/sets";
 
-val induct_meth =
-  P.and_list' (Scan.repeat (unless_more_args def_inst)) --
-    (arbitrary -- taking -- Scan.option induct_rule) >>
-  (fn (insts, ((arbitrary, taking), opt_rule)) => fn ctxt =>
-    RAW_METHOD_CASES (fn facts =>
-      Seq.DETERM (HEADGOAL (induct_tac ctxt insts arbitrary taking opt_rule facts))));
+val induct_setup =
+  Method.setup @{binding induct}
+    (P.and_list' (Scan.repeat (unless_more_args def_inst)) --
+      (arbitrary -- taking -- Scan.option induct_rule) >>
+      (fn (insts, ((arbitrary, taking), opt_rule)) => fn ctxt =>
+        RAW_METHOD_CASES (fn facts =>
+          Seq.DETERM (HEADGOAL (induct_tac ctxt insts arbitrary taking opt_rule facts)))))
+    "induction on types or predicates/sets";
 
-val coinduct_meth =
-  Scan.repeat (unless_more_args inst) -- taking -- Scan.option coinduct_rule >>
-  (fn ((insts, taking), opt_rule) => fn ctxt =>
-    RAW_METHOD_CASES (fn facts =>
-      Seq.DETERM (HEADGOAL (coinduct_tac ctxt insts taking opt_rule facts))));
+val coinduct_setup =
+  Method.setup @{binding coinduct}
+    (Scan.repeat (unless_more_args inst) -- taking -- Scan.option coinduct_rule >>
+      (fn ((insts, taking), opt_rule) => fn ctxt =>
+        RAW_METHOD_CASES (fn facts =>
+          Seq.DETERM (HEADGOAL (coinduct_tac ctxt insts taking opt_rule facts)))))
+    "coinduction on types or predicates/sets";
 
 end;
 
@@ -759,10 +764,6 @@
 
 (** theory setup **)
 
-val setup =
-  attrib_setup #>
-  Method.setup @{binding cases} cases_meth "case analysis on types or predicates/sets" #>
-  Method.setup @{binding induct} induct_meth "induction on types or predicates/sets" #>
-  Method.setup @{binding coinduct} coinduct_meth "coinduction on types or predicates/sets";
+val setup = attrib_setup #> cases_setup  #> induct_setup #> coinduct_setup;
 
 end;
--- a/src/Tools/nbe.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/Tools/nbe.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -277,14 +277,12 @@
       in
         s
         |> tracing (fn s => "\n--- code to be evaluated:\n" ^ s)
-        |> ML_Context.evaluate ctxt
-            (Output.tracing o enclose "\n---compiler echo:\n" "\n---\n",
-            Output.tracing o enclose "\n--- compiler echo (with error):\n" "\n---\n")
-            (!trace) univs_cookie
+        |> ML_Context.evaluate ctxt (!trace) univs_cookie
         |> (fn f => f deps_vals)
         |> (fn univs => cs ~~ univs)
       end;
 
+
 (* preparing function equations *)
 
 fun eqns_of_stmt (_, Code_Thingol.Fun (_, (_, []))) =
--- a/src/ZF/Constructible/L_axioms.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/ZF/Constructible/L_axioms.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -99,7 +99,7 @@
   apply (rule L_nat)
   done
 
-interpretation L: M_trivial L by (rule M_trivial_L)
+interpretation L?: M_trivial L by (rule M_trivial_L)
 
 (* Replaces the following declarations...
 lemmas rall_abs = M_trivial.rall_abs [OF M_trivial_L]
--- a/src/ZF/Constructible/Rec_Separation.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/ZF/Constructible/Rec_Separation.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -185,7 +185,7 @@
 theorem M_trancl_L: "PROP M_trancl(L)"
 by (rule M_trancl.intro [OF M_basic_L M_trancl_axioms_L])
 
-interpretation L: M_trancl L by (rule M_trancl_L)
+interpretation L?: M_trancl L by (rule M_trancl_L)
 
 
 subsection{*@{term L} is Closed Under the Operator @{term list}*}
@@ -372,7 +372,7 @@
   apply (rule M_datatypes_axioms_L) 
   done
 
-interpretation L: M_datatypes L by (rule M_datatypes_L)
+interpretation L?: M_datatypes L by (rule M_datatypes_L)
 
 
 subsection{*@{term L} is Closed Under the Operator @{term eclose}*}
@@ -435,7 +435,7 @@
   apply (rule M_eclose_axioms_L)
   done
 
-interpretation L: M_eclose L by (rule M_eclose_L)
+interpretation L?: M_eclose L by (rule M_eclose_L)
 
 
 end
--- a/src/ZF/Constructible/Separation.thy	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/ZF/Constructible/Separation.thy	Sat Mar 28 00:13:01 2009 +0100
@@ -305,7 +305,7 @@
 theorem M_basic_L: "PROP M_basic(L)"
 by (rule M_basic.intro [OF M_trivial_L M_basic_axioms_L])
 
-interpretation L: M_basic L by (rule M_basic_L)
+interpretation L?: M_basic L by (rule M_basic_L)
 
 
 end
--- a/src/ZF/Tools/ind_cases.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/ZF/Tools/ind_cases.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -57,14 +57,14 @@
 
 (* ind_cases method *)
 
-val mk_cases_meth = Scan.lift (Scan.repeat1 Args.name_source) >>
-  (fn props => fn ctxt =>
-    props
-    |> map (smart_cases (ProofContext.theory_of ctxt) (local_simpset_of ctxt) (Syntax.read_prop ctxt))
-    |> Method.erule 0);
-
 val setup =
-  Method.setup @{binding "ind_cases"} mk_cases_meth "dynamic case analysis on sets";
+  Method.setup @{binding "ind_cases"}
+    (Scan.lift (Scan.repeat1 Args.name_source) >>
+      (fn props => fn ctxt =>
+        props
+        |> map (smart_cases (ProofContext.theory_of ctxt) (local_simpset_of ctxt) (Syntax.read_prop ctxt))
+        |> Method.erule 0))
+    "dynamic case analysis on sets";
 
 
 (* outer syntax *)
--- a/src/ZF/Tools/inductive_package.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/ZF/Tools/inductive_package.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -16,7 +16,6 @@
     dom_subset : thm,                  (*inclusion of recursive set in dom*)
     intrs      : thm list,             (*introduction rules*)
     elim       : thm,                  (*case analysis theorem*)
-    mk_cases   : string -> thm,        (*generates case theorems*)
     induct     : thm,                  (*main induction rule*)
     mutual_induct : thm};              (*mutual induction rule*)
 
@@ -231,12 +230,12 @@
      if !Ind_Syntax.trace then print_tac "The type-checking subgoal:"
      else all_tac,
      REPEAT (FIRSTGOAL (        dresolve_tac rec_typechecks
-                        ORELSE' eresolve_tac (asm_rl::@{thm PartE}::@{thm SigmaE2}::
+                        ORELSE' eresolve_tac (asm_rl :: @{thm PartE} :: @{thm SigmaE2} ::
                                               type_elims)
                         ORELSE' hyp_subst_tac)),
      if !Ind_Syntax.trace then print_tac "The subgoal after monos, type_elims:"
      else all_tac,
-     DEPTH_SOLVE (swap_res_tac (@{thm SigmaI}::@{thm subsetI}::type_intrs) 1)];
+     DEPTH_SOLVE (swap_res_tac (@{thm SigmaI} :: @{thm subsetI} :: type_intrs) 1)];
 
   (*combines disjI1 and disjI2 to get the corresponding nested disjunct...*)
   val mk_disj_rls = BalancedTree.accesses
@@ -275,9 +274,6 @@
       (basic_elim_tac THEN ALLGOALS (asm_full_simp_tac ss) THEN basic_elim_tac)
       (Thm.assume A RS elim)
       |> Drule.standard';
-  fun mk_cases a = make_cases (*delayed evaluation of body!*)
-    (simpset ())
-    let val thy = Thm.theory_of_thm elim in cterm_of thy (Syntax.read_prop_global thy a) end;
 
   fun induction_rules raw_induct thy =
    let
@@ -548,7 +544,6 @@
        dom_subset = dom_subset',
        intrs = intrs'',
        elim = elim',
-       mk_cases = mk_cases,
        induct = induct,
        mutual_induct = mutual_induct})
   end;
--- a/src/ZF/Tools/typechk.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/ZF/Tools/typechk.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -109,11 +109,13 @@
 
 (* concrete syntax *)
 
-val typecheck_meth =
-  Method.sections
-    [Args.add -- Args.colon >> K (I, TC_add),
-     Args.del -- Args.colon >> K (I, TC_del)]
-  >> (K (fn ctxt => SIMPLE_METHOD (CHANGED (typecheck_tac ctxt))));
+val typecheck_setup =
+  Method.setup @{binding typecheck}
+    (Method.sections
+      [Args.add -- Args.colon >> K (I, TC_add),
+       Args.del -- Args.colon >> K (I, TC_del)]
+      >> (K (fn ctxt => SIMPLE_METHOD (CHANGED (typecheck_tac ctxt)))))
+    "ZF type-checking";
 
 val _ =
   OuterSyntax.improper_command "print_tcset" "print context of ZF typecheck" OuterKeyword.diag
@@ -125,7 +127,7 @@
 
 val setup =
   Attrib.setup @{binding TC} (Attrib.add_del TC_add TC_del) "declaration of type-checking rule" #>
-  Method.setup @{binding typecheck} typecheck_meth "ZF type-checking" #>
+  typecheck_setup #>
   Simplifier.map_simpset (fn ss => ss setSolver type_solver);
 
 end;
--- a/src/ZF/int_arith.ML	Sat Mar 28 00:11:02 2009 +0100
+++ b/src/ZF/int_arith.ML	Sat Mar 28 00:13:01 2009 +0100
@@ -145,7 +145,7 @@
   val numeral_simp_ss = ZF_ss addsimps add_0s @ bin_simps @ tc_rules @ intifys
   fun numeral_simp_tac ss =
     ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss))
-    THEN ALLGOALS (SIMPSET' (fn simpset => asm_simp_tac (Simplifier.inherit_context ss simpset)))
+    THEN ALLGOALS (asm_simp_tac (local_simpset_of (Simplifier.the_context ss)))
   val simplify_meta_eq  = ArithData.simplify_meta_eq (add_0s @ mult_1s)
   end;